src/HOL/Import/HOLLight/HOLLight.thy
author blanchet
Wed, 28 Apr 2010 18:11:11 +0200
changeset 36550 f8da913b6c3a
parent 35416 d8d7d1b785af
child 43786 fea3ed6951e3
permissions -rw-r--r--
make sure short theorem names are preferred to composite ones in Sledgehammer; this code used to work at some point
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
     1
(* AUTOMATICALLY GENERATED, DO NOT EDIT! *)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
     2
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
     3
theory HOLLight = "../HOLLightCompat" + "../HOL4Syntax":
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
     4
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
     5
;setup_theory hollight
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
     6
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
     7
consts
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
     8
  "_FALSITY_" :: "bool" ("'_FALSITY'_")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
     9
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    10
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    11
  "_FALSITY__def": "_FALSITY_ == False"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    12
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    13
lemma DEF__FALSITY_: "_FALSITY_ = False"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    14
  by (import hollight DEF__FALSITY_)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    15
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    16
lemma CONJ_ACI: "((p::bool) & (q::bool)) = (q & p) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    17
((p & q) & (r::bool)) = (p & q & r) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    18
(p & q & r) = (q & p & r) & (p & p) = p & (p & p & q) = (p & q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    19
  by (import hollight CONJ_ACI)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    20
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    21
lemma DISJ_ACI: "((p::bool) | (q::bool)) = (q | p) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    22
((p | q) | (r::bool)) = (p | q | r) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    23
(p | q | r) = (q | p | r) & (p | p) = p & (p | p | q) = (p | q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    24
  by (import hollight DISJ_ACI)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    25
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    26
lemma EQ_CLAUSES: "ALL t::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    27
   (True = t) = t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    28
   (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    29
  by (import hollight EQ_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    30
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    31
lemma NOT_CLAUSES_WEAK: "(~ True) = False & (~ False) = True"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    32
  by (import hollight NOT_CLAUSES_WEAK)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    33
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    34
lemma AND_CLAUSES: "ALL t::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    35
   (True & t) = t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    36
   (t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    37
  by (import hollight AND_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    38
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    39
lemma OR_CLAUSES: "ALL t::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    40
   (True | t) = True &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    41
   (t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    42
  by (import hollight OR_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    43
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    44
lemma IMP_CLAUSES: "ALL t::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    45
   (True --> t) = t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    46
   (t --> True) = True &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    47
   (False --> t) = True & (t --> t) = True & (t --> False) = (~ t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    48
  by (import hollight IMP_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    49
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    50
lemma IMP_EQ_CLAUSE: "((x::'q_864::type) = x --> (p::bool)) = p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    51
  by (import hollight IMP_EQ_CLAUSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    52
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    53
lemma SWAP_FORALL_THM: "ALL P::'A::type => 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    54
   (ALL x::'A::type. All (P x)) = (ALL (y::'B::type) x::'A::type. P x y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    55
  by (import hollight SWAP_FORALL_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    56
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    57
lemma SWAP_EXISTS_THM: "ALL P::'A::type => 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    58
   (EX x::'A::type. Ex (P x)) = (EX (x::'B::type) xa::'A::type. P xa x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    59
  by (import hollight SWAP_EXISTS_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    60
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    61
lemma TRIV_EXISTS_AND_THM: "ALL (P::bool) Q::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    62
   (EX x::'A::type. P & Q) = ((EX x::'A::type. P) & (EX x::'A::type. Q))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    63
  by (import hollight TRIV_EXISTS_AND_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    64
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    65
lemma TRIV_AND_EXISTS_THM: "ALL (P::bool) Q::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    66
   ((EX x::'A::type. P) & (EX x::'A::type. Q)) = (EX x::'A::type. P & Q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    67
  by (import hollight TRIV_AND_EXISTS_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    68
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    69
lemma TRIV_FORALL_OR_THM: "ALL (P::bool) Q::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    70
   (ALL x::'A::type. P | Q) = ((ALL x::'A::type. P) | (ALL x::'A::type. Q))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    71
  by (import hollight TRIV_FORALL_OR_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    72
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    73
lemma TRIV_OR_FORALL_THM: "ALL (P::bool) Q::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    74
   ((ALL x::'A::type. P) | (ALL x::'A::type. Q)) = (ALL x::'A::type. P | Q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    75
  by (import hollight TRIV_OR_FORALL_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    76
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    77
lemma TRIV_FORALL_IMP_THM: "ALL (P::bool) Q::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    78
   (ALL x::'A::type. P --> Q) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    79
   ((EX x::'A::type. P) --> (ALL x::'A::type. Q))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    80
  by (import hollight TRIV_FORALL_IMP_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    81
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    82
lemma TRIV_EXISTS_IMP_THM: "ALL (P::bool) Q::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    83
   (EX x::'A::type. P --> Q) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    84
   ((ALL x::'A::type. P) --> (EX x::'A::type. Q))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    85
  by (import hollight TRIV_EXISTS_IMP_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    86
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    87
lemma EXISTS_UNIQUE_ALT: "ALL P::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    88
   Ex1 P = (EX x::'A::type. ALL y::'A::type. P y = (x = y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    89
  by (import hollight EXISTS_UNIQUE_ALT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    90
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    91
lemma SELECT_UNIQUE: "ALL (P::'A::type => bool) x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    92
   (ALL y::'A::type. P y = (y = x)) --> Eps P = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    93
  by (import hollight SELECT_UNIQUE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    94
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    95
lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    96
  by (import hollight EXCLUDED_MIDDLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    97
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
    98
definition COND :: "bool => 'A => 'A => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
    99
  "COND ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   100
%(t::bool) (t1::'A::type) t2::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   101
   SOME x::'A::type. (t = True --> x = t1) & (t = False --> x = t2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   102
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   103
lemma DEF_COND: "COND =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   104
(%(t::bool) (t1::'A::type) t2::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   105
    SOME x::'A::type. (t = True --> x = t1) & (t = False --> x = t2))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   106
  by (import hollight DEF_COND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   107
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   108
lemma COND_CLAUSES: "ALL (x::'A::type) xa::'A::type. COND True x xa = x & COND False x xa = xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   109
  by (import hollight COND_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   110
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   111
lemma COND_EXPAND: "ALL (b::bool) (t1::bool) t2::bool. COND b t1 t2 = ((~ b | t1) & (b | t2))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   112
  by (import hollight COND_EXPAND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   113
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   114
lemma COND_ID: "ALL (b::bool) t::'A::type. COND b t t = t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   115
  by (import hollight COND_ID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   116
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   117
lemma COND_RAND: "ALL (b::bool) (f::'A::type => 'B::type) (x::'A::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   118
   f (COND b x y) = COND b (f x) (f y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   119
  by (import hollight COND_RAND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   120
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   121
lemma COND_RATOR: "ALL (b::bool) (f::'A::type => 'B::type) (g::'A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   122
   x::'A::type. COND b f g x = COND b (f x) (g x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   123
  by (import hollight COND_RATOR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   124
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   125
lemma COND_ABS: "ALL (b::bool) (f::'A::type => 'B::type) g::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   126
   (%x::'A::type. COND b (f x) (g x)) = COND b f g"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   127
  by (import hollight COND_ABS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   128
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   129
lemma MONO_COND: "((A::bool) --> (B::bool)) & ((C::bool) --> (D::bool)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   130
COND (b::bool) A C --> COND b B D"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   131
  by (import hollight MONO_COND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   132
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   133
lemma COND_ELIM_THM: "(P::'A::type => bool) (COND (c::bool) (x::'A::type) (y::'A::type)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   134
((c --> P x) & (~ c --> P y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   135
  by (import hollight COND_ELIM_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   136
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   137
lemma SKOLEM_THM: "ALL P::'A::type => 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   138
   (ALL x::'A::type. Ex (P x)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   139
   (EX x::'A::type => 'B::type. ALL xa::'A::type. P xa (x xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   140
  by (import hollight SKOLEM_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   141
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   142
lemma UNIQUE_SKOLEM_ALT: "ALL P::'A::type => 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   143
   (ALL x::'A::type. Ex1 (P x)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   144
   (EX f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   145
       ALL (x::'A::type) y::'B::type. P x y = (f x = y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   146
  by (import hollight UNIQUE_SKOLEM_ALT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   147
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   148
lemma COND_EQ_CLAUSE: "COND ((x::'q_3000::type) = x) (y::'q_2993::type) (z::'q_2993::type) = y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   149
  by (import hollight COND_EQ_CLAUSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   150
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   151
lemma o_ASSOC: "ALL (f::'C::type => 'D::type) (g::'B::type => 'C::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   152
   h::'A::type => 'B::type. f o (g o h) = f o g o h"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   153
  by (import hollight o_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   154
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   155
lemma I_O_ID: "ALL f::'A::type => 'B::type. id o f = f & f o id = f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   156
  by (import hollight I_O_ID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   157
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   158
lemma EXISTS_ONE_REP: "EX x::bool. x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   159
  by (import hollight EXISTS_ONE_REP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   160
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   161
lemma one_axiom: "ALL f::'A::type => unit. All (op = f)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   162
  by (import hollight one_axiom)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   163
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   164
lemma one_RECURSION: "ALL e::'A::type. EX x::unit => 'A::type. x () = e"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   165
  by (import hollight one_RECURSION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   166
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   167
lemma one_Axiom: "ALL e::'A::type. EX! fn::unit => 'A::type. fn () = e"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   168
  by (import hollight one_Axiom)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   169
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   170
lemma th_cond: "(P::'A::type => bool => bool) (COND (b::bool) (x::'A::type) (y::'A::type))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   171
 b =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   172
(b & P x True | ~ b & P y False)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   173
  by (import hollight th_cond)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   174
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   175
definition LET_END :: "'A => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   176
  "LET_END == %t::'A::type. t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   177
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   178
lemma DEF_LET_END: "LET_END = (%t::'A::type. t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   179
  by (import hollight DEF_LET_END)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   180
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   181
definition GABS :: "('A => bool) => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   182
  "(op ==::(('A::type => bool) => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   183
        => (('A::type => bool) => 'A::type) => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   184
 (GABS::('A::type => bool) => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   185
 (Eps::('A::type => bool) => 'A::type)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   186
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   187
lemma DEF_GABS: "(op =::(('A::type => bool) => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   188
       => (('A::type => bool) => 'A::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   189
 (GABS::('A::type => bool) => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   190
 (Eps::('A::type => bool) => 'A::type)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   191
  by (import hollight DEF_GABS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   192
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   193
definition GEQ :: "'A => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   194
  "(op ==::('A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   195
        => ('A::type => 'A::type => bool) => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   196
 (GEQ::'A::type => 'A::type => bool) (op =::'A::type => 'A::type => bool)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   197
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   198
lemma DEF_GEQ: "(op =::('A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   199
       => ('A::type => 'A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   200
 (GEQ::'A::type => 'A::type => bool) (op =::'A::type => 'A::type => bool)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   201
  by (import hollight DEF_GEQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   202
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   203
lemma PAIR_EXISTS_THM: "EX (x::'A::type => 'B::type => bool) (a::'A::type) b::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   204
   x = Pair_Rep a b"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   205
  by (import hollight PAIR_EXISTS_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   206
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   207
definition CURRY :: "('A * 'B => 'C) => 'A => 'B => 'C" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   208
  "CURRY ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   209
%(u::'A::type * 'B::type => 'C::type) (ua::'A::type) ub::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   210
   u (ua, ub)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   211
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   212
lemma DEF_CURRY: "CURRY =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   213
(%(u::'A::type * 'B::type => 'C::type) (ua::'A::type) ub::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   214
    u (ua, ub))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   215
  by (import hollight DEF_CURRY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   216
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   217
definition UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   218
  "UNCURRY ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   219
%(u::'A::type => 'B::type => 'C::type) ua::'A::type * 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   220
   u (fst ua) (snd ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   221
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   222
lemma DEF_UNCURRY: "UNCURRY =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   223
(%(u::'A::type => 'B::type => 'C::type) ua::'A::type * 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   224
    u (fst ua) (snd ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   225
  by (import hollight DEF_UNCURRY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   226
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   227
definition PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   228
  "PASSOC ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   229
%(u::('A::type * 'B::type) * 'C::type => 'D::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   230
   ua::'A::type * 'B::type * 'C::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   231
   u ((fst ua, fst (snd ua)), snd (snd ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   232
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   233
lemma DEF_PASSOC: "PASSOC =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   234
(%(u::('A::type * 'B::type) * 'C::type => 'D::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   235
    ua::'A::type * 'B::type * 'C::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   236
    u ((fst ua, fst (snd ua)), snd (snd ua)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   237
  by (import hollight DEF_PASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   238
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   239
lemma num_Axiom: "ALL (e::'A::type) f::'A::type => nat => 'A::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   240
   EX! fn::nat => 'A::type. fn 0 = e & (ALL n::nat. fn (Suc n) = f (fn n) n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   241
  by (import hollight num_Axiom)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   242
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   243
lemma ADD_CLAUSES: "(ALL x::nat. 0 + x = x) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   244
(ALL x::nat. x + 0 = x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   245
(ALL (x::nat) xa::nat. Suc x + xa = Suc (x + xa)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   246
(ALL (x::nat) xa::nat. x + Suc xa = Suc (x + xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   247
  by (import hollight ADD_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   248
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   249
lemma ADD_AC: "(m::nat) + (n::nat) = n + m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   250
m + n + (p::nat) = m + (n + p) & m + (n + p) = n + (m + p)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   251
  by (import hollight ADD_AC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   252
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   253
lemma EQ_ADD_LCANCEL_0: "ALL (m::nat) n::nat. (m + n = m) = (n = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   254
  by (import hollight EQ_ADD_LCANCEL_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   255
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   256
lemma EQ_ADD_RCANCEL_0: "ALL (x::nat) xa::nat. (x + xa = xa) = (x = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   257
  by (import hollight EQ_ADD_RCANCEL_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   258
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   259
lemma ONE: "NUMERAL_BIT1 0 = Suc 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   260
  by (import hollight ONE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   261
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   262
lemma TWO: "NUMERAL_BIT0 (NUMERAL_BIT1 0) = Suc (NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   263
  by (import hollight TWO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   264
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   265
lemma ADD1: "ALL x::nat. Suc x = x + NUMERAL_BIT1 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   266
  by (import hollight ADD1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   267
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   268
lemma MULT_CLAUSES: "(ALL x::nat. 0 * x = 0) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   269
(ALL x::nat. x * 0 = 0) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   270
(ALL x::nat. NUMERAL_BIT1 0 * x = x) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   271
(ALL x::nat. x * NUMERAL_BIT1 0 = x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   272
(ALL (x::nat) xa::nat. Suc x * xa = x * xa + xa) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   273
(ALL (x::nat) xa::nat. x * Suc xa = x + x * xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   274
  by (import hollight MULT_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   275
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   276
lemma MULT_AC: "(m::nat) * (n::nat) = n * m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   277
m * n * (p::nat) = m * (n * p) & m * (n * p) = n * (m * p)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   278
  by (import hollight MULT_AC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   279
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   280
lemma MULT_2: "ALL n::nat. NUMERAL_BIT0 (NUMERAL_BIT1 0) * n = n + n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   281
  by (import hollight MULT_2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   282
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   283
lemma MULT_EQ_1: "ALL (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   284
   (m * n = NUMERAL_BIT1 0) = (m = NUMERAL_BIT1 0 & n = NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   285
  by (import hollight MULT_EQ_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   286
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   287
definition EXP :: "nat => nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   288
  "EXP ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   289
SOME EXP::nat => nat => nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   290
   (ALL m::nat. EXP m 0 = NUMERAL_BIT1 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   291
   (ALL (m::nat) n::nat. EXP m (Suc n) = m * EXP m n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   292
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   293
lemma DEF_EXP: "EXP =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   294
(SOME EXP::nat => nat => nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   295
    (ALL m::nat. EXP m 0 = NUMERAL_BIT1 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   296
    (ALL (m::nat) n::nat. EXP m (Suc n) = m * EXP m n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   297
  by (import hollight DEF_EXP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   298
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   299
lemma EXP_EQ_0: "ALL (m::nat) n::nat. (EXP m n = 0) = (m = 0 & n ~= 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   300
  by (import hollight EXP_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   301
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   302
lemma EXP_ADD: "ALL (m::nat) (n::nat) p::nat. EXP m (n + p) = EXP m n * EXP m p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   303
  by (import hollight EXP_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   304
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   305
lemma EXP_ONE: "ALL n::nat. EXP (NUMERAL_BIT1 0) n = NUMERAL_BIT1 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   306
  by (import hollight EXP_ONE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   307
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   308
lemma EXP_1: "ALL x::nat. EXP x (NUMERAL_BIT1 0) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   309
  by (import hollight EXP_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   310
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   311
lemma EXP_2: "ALL x::nat. EXP x (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = x * x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   312
  by (import hollight EXP_2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   313
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   314
lemma MULT_EXP: "ALL (p::nat) (m::nat) n::nat. EXP (m * n) p = EXP m p * EXP n p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   315
  by (import hollight MULT_EXP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   316
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   317
lemma EXP_MULT: "ALL (m::nat) (n::nat) p::nat. EXP m (n * p) = EXP (EXP m n) p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   318
  by (import hollight EXP_MULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   319
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   320
consts
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   321
  "<=" :: "nat => nat => bool" ("<=")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   322
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   323
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   324
  "<=_def": "<= ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   325
SOME u::nat => nat => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   326
   (ALL m::nat. u m 0 = (m = 0)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   327
   (ALL (m::nat) n::nat. u m (Suc n) = (m = Suc n | u m n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   328
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   329
lemma DEF__lessthan__equal_: "<= =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   330
(SOME u::nat => nat => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   331
    (ALL m::nat. u m 0 = (m = 0)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   332
    (ALL (m::nat) n::nat. u m (Suc n) = (m = Suc n | u m n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   333
  by (import hollight DEF__lessthan__equal_)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   334
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   335
consts
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   336
  "<" :: "nat => nat => bool" ("<")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   337
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   338
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   339
  "<_def": "< ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   340
SOME u::nat => nat => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   341
   (ALL m::nat. u m 0 = False) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   342
   (ALL (m::nat) n::nat. u m (Suc n) = (m = n | u m n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   343
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   344
lemma DEF__lessthan_: "< =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   345
(SOME u::nat => nat => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   346
    (ALL m::nat. u m 0 = False) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   347
    (ALL (m::nat) n::nat. u m (Suc n) = (m = n | u m n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   348
  by (import hollight DEF__lessthan_)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   349
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   350
consts
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   351
  ">=" :: "nat => nat => bool" (">=")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   352
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   353
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   354
  ">=_def": ">= == %(u::nat) ua::nat. <= ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   355
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   356
lemma DEF__greaterthan__equal_: ">= = (%(u::nat) ua::nat. <= ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   357
  by (import hollight DEF__greaterthan__equal_)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   358
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   359
consts
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   360
  ">" :: "nat => nat => bool" (">")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   361
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   362
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   363
  ">_def": "> == %(u::nat) ua::nat. < ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   364
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   365
lemma DEF__greaterthan_: "> = (%(u::nat) ua::nat. < ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   366
  by (import hollight DEF__greaterthan_)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   367
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   368
lemma LE_SUC_LT: "ALL (m::nat) n::nat. <= (Suc m) n = < m n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   369
  by (import hollight LE_SUC_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   370
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   371
lemma LT_SUC_LE: "ALL (m::nat) n::nat. < m (Suc n) = <= m n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   372
  by (import hollight LT_SUC_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   373
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   374
lemma LE_SUC: "ALL (x::nat) xa::nat. <= (Suc x) (Suc xa) = <= x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   375
  by (import hollight LE_SUC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   376
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   377
lemma LT_SUC: "ALL (x::nat) xa::nat. < (Suc x) (Suc xa) = < x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   378
  by (import hollight LT_SUC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   379
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   380
lemma LE_0: "All (<= 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   381
  by (import hollight LE_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   382
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   383
lemma LT_0: "ALL x::nat. < 0 (Suc x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   384
  by (import hollight LT_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   385
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   386
lemma LE_REFL: "ALL n::nat. <= n n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   387
  by (import hollight LE_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   388
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   389
lemma LT_REFL: "ALL n::nat. ~ < n n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   390
  by (import hollight LT_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   391
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   392
lemma LE_ANTISYM: "ALL (m::nat) n::nat. (<= m n & <= n m) = (m = n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   393
  by (import hollight LE_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   394
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   395
lemma LT_ANTISYM: "ALL (m::nat) n::nat. ~ (< m n & < n m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   396
  by (import hollight LT_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   397
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   398
lemma LET_ANTISYM: "ALL (m::nat) n::nat. ~ (<= m n & < n m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   399
  by (import hollight LET_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   400
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   401
lemma LTE_ANTISYM: "ALL (x::nat) xa::nat. ~ (< x xa & <= xa x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   402
  by (import hollight LTE_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   403
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   404
lemma LE_TRANS: "ALL (m::nat) (n::nat) p::nat. <= m n & <= n p --> <= m p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   405
  by (import hollight LE_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   406
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   407
lemma LT_TRANS: "ALL (m::nat) (n::nat) p::nat. < m n & < n p --> < m p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   408
  by (import hollight LT_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   409
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   410
lemma LET_TRANS: "ALL (m::nat) (n::nat) p::nat. <= m n & < n p --> < m p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   411
  by (import hollight LET_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   412
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   413
lemma LTE_TRANS: "ALL (m::nat) (n::nat) p::nat. < m n & <= n p --> < m p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   414
  by (import hollight LTE_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   415
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   416
lemma LE_CASES: "ALL (m::nat) n::nat. <= m n | <= n m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   417
  by (import hollight LE_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   418
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   419
lemma LT_CASES: "ALL (m::nat) n::nat. < m n | < n m | m = n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   420
  by (import hollight LT_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   421
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   422
lemma LET_CASES: "ALL (m::nat) n::nat. <= m n | < n m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   423
  by (import hollight LET_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   424
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   425
lemma LTE_CASES: "ALL (x::nat) xa::nat. < x xa | <= xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   426
  by (import hollight LTE_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   427
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   428
lemma LT_NZ: "ALL n::nat. < 0 n = (n ~= 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   429
  by (import hollight LT_NZ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   430
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   431
lemma LE_LT: "ALL (m::nat) n::nat. <= m n = (< m n | m = n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   432
  by (import hollight LE_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   433
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   434
lemma LT_LE: "ALL (x::nat) xa::nat. < x xa = (<= x xa & x ~= xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   435
  by (import hollight LT_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   436
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   437
lemma NOT_LE: "ALL (m::nat) n::nat. (~ <= m n) = < n m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   438
  by (import hollight NOT_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   439
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   440
lemma NOT_LT: "ALL (m::nat) n::nat. (~ < m n) = <= n m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   441
  by (import hollight NOT_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   442
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   443
lemma LT_IMP_LE: "ALL (x::nat) xa::nat. < x xa --> <= x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   444
  by (import hollight LT_IMP_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   445
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   446
lemma EQ_IMP_LE: "ALL (m::nat) n::nat. m = n --> <= m n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   447
  by (import hollight EQ_IMP_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   448
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   449
lemma LE_EXISTS: "ALL (m::nat) n::nat. <= m n = (EX d::nat. n = m + d)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   450
  by (import hollight LE_EXISTS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   451
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   452
lemma LT_EXISTS: "ALL (m::nat) n::nat. < m n = (EX d::nat. n = m + Suc d)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   453
  by (import hollight LT_EXISTS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   454
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   455
lemma LE_ADD: "ALL (m::nat) n::nat. <= m (m + n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   456
  by (import hollight LE_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   457
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   458
lemma LE_ADDR: "ALL (x::nat) xa::nat. <= xa (x + xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   459
  by (import hollight LE_ADDR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   460
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   461
lemma LT_ADD: "ALL (m::nat) n::nat. < m (m + n) = < 0 n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   462
  by (import hollight LT_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   463
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   464
lemma LT_ADDR: "ALL (x::nat) xa::nat. < xa (x + xa) = < 0 x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   465
  by (import hollight LT_ADDR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   466
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   467
lemma LE_ADD_LCANCEL: "ALL (x::nat) (xa::nat) xb::nat. <= (x + xa) (x + xb) = <= xa xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   468
  by (import hollight LE_ADD_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   469
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   470
lemma LE_ADD_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. <= (x + xb) (xa + xb) = <= x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   471
  by (import hollight LE_ADD_RCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   472
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   473
lemma LT_ADD_LCANCEL: "ALL (x::nat) (xa::nat) xb::nat. < (x + xa) (x + xb) = < xa xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   474
  by (import hollight LT_ADD_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   475
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   476
lemma LT_ADD_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. < (x + xb) (xa + xb) = < x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   477
  by (import hollight LT_ADD_RCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   478
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   479
lemma LE_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   480
   <= m p & <= n q --> <= (m + n) (p + q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   481
  by (import hollight LE_ADD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   482
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   483
lemma LET_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. <= m p & < n q --> < (m + n) (p + q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   484
  by (import hollight LET_ADD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   485
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   486
lemma LTE_ADD2: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   487
   < x xb & <= xa xc --> < (x + xa) (xb + xc)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   488
  by (import hollight LTE_ADD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   489
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   490
lemma LT_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. < m p & < n q --> < (m + n) (p + q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   491
  by (import hollight LT_ADD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   492
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   493
lemma LT_MULT: "ALL (m::nat) n::nat. < 0 (m * n) = (< 0 m & < 0 n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   494
  by (import hollight LT_MULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   495
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   496
lemma LE_MULT2: "ALL (m::nat) (n::nat) (p::nat) q::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   497
   <= m n & <= p q --> <= (m * p) (n * q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   498
  by (import hollight LE_MULT2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   499
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   500
lemma LT_LMULT: "ALL (m::nat) (n::nat) p::nat. m ~= 0 & < n p --> < (m * n) (m * p)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   501
  by (import hollight LT_LMULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   502
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   503
lemma LE_MULT_LCANCEL: "ALL (m::nat) (n::nat) p::nat. <= (m * n) (m * p) = (m = 0 | <= n p)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   504
  by (import hollight LE_MULT_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   505
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   506
lemma LE_MULT_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. <= (x * xb) (xa * xb) = (<= x xa | xb = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   507
  by (import hollight LE_MULT_RCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   508
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   509
lemma LT_MULT_LCANCEL: "ALL (m::nat) (n::nat) p::nat. < (m * n) (m * p) = (m ~= 0 & < n p)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   510
  by (import hollight LT_MULT_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   511
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   512
lemma LT_MULT_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. < (x * xb) (xa * xb) = (< x xa & xb ~= 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   513
  by (import hollight LT_MULT_RCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   514
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   515
lemma LT_MULT2: "ALL (m::nat) (n::nat) (p::nat) q::nat. < m n & < p q --> < (m * p) (n * q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   516
  by (import hollight LT_MULT2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   517
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   518
lemma LE_SQUARE_REFL: "ALL n::nat. <= n (n * n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   519
  by (import hollight LE_SQUARE_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   520
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   521
lemma WLOG_LE: "(ALL (m::nat) n::nat. (P::nat => nat => bool) m n = P n m) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   522
(ALL (m::nat) n::nat. <= m n --> P m n) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   523
(ALL m::nat. All (P m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   524
  by (import hollight WLOG_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   525
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   526
lemma WLOG_LT: "(ALL m::nat. (P::nat => nat => bool) m m) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   527
(ALL (m::nat) n::nat. P m n = P n m) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   528
(ALL (m::nat) n::nat. < m n --> P m n) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   529
(ALL m::nat. All (P m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   530
  by (import hollight WLOG_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   531
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   532
lemma num_WF: "ALL P::nat => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   533
   (ALL n::nat. (ALL m::nat. < m n --> P m) --> P n) --> All P"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   534
  by (import hollight num_WF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   535
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   536
lemma num_WOP: "ALL P::nat => bool. Ex P = (EX n::nat. P n & (ALL m::nat. < m n --> ~ P m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   537
  by (import hollight num_WOP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   538
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   539
lemma num_MAX: "ALL P::nat => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   540
   (Ex P & (EX M::nat. ALL x::nat. P x --> <= x M)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   541
   (EX m::nat. P m & (ALL x::nat. P x --> <= x m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   542
  by (import hollight num_MAX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   543
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   544
definition EVEN :: "nat => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   545
  "EVEN ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   546
SOME EVEN::nat => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   547
   EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   548
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   549
lemma DEF_EVEN: "EVEN =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   550
(SOME EVEN::nat => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   551
    EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   552
  by (import hollight DEF_EVEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   553
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   554
definition ODD :: "nat => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   555
  "ODD ==
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   556
SOME ODD::nat => bool. ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   557
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   558
lemma DEF_ODD: "ODD =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   559
(SOME ODD::nat => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   560
    ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   561
  by (import hollight DEF_ODD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   562
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   563
lemma NOT_EVEN: "ALL n::nat. (~ EVEN n) = ODD n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   564
  by (import hollight NOT_EVEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   565
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   566
lemma NOT_ODD: "ALL n::nat. (~ ODD n) = EVEN n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   567
  by (import hollight NOT_ODD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   568
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   569
lemma EVEN_OR_ODD: "ALL n::nat. EVEN n | ODD n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   570
  by (import hollight EVEN_OR_ODD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   571
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   572
lemma EVEN_AND_ODD: "ALL x::nat. ~ (EVEN x & ODD x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   573
  by (import hollight EVEN_AND_ODD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   574
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   575
lemma EVEN_ADD: "ALL (m::nat) n::nat. EVEN (m + n) = (EVEN m = EVEN n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   576
  by (import hollight EVEN_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   577
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   578
lemma EVEN_MULT: "ALL (m::nat) n::nat. EVEN (m * n) = (EVEN m | EVEN n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   579
  by (import hollight EVEN_MULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   580
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   581
lemma EVEN_EXP: "ALL (m::nat) n::nat. EVEN (EXP m n) = (EVEN m & n ~= 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   582
  by (import hollight EVEN_EXP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   583
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   584
lemma ODD_ADD: "ALL (m::nat) n::nat. ODD (m + n) = (ODD m ~= ODD n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   585
  by (import hollight ODD_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   586
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   587
lemma ODD_MULT: "ALL (m::nat) n::nat. ODD (m * n) = (ODD m & ODD n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   588
  by (import hollight ODD_MULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   589
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   590
lemma ODD_EXP: "ALL (m::nat) n::nat. ODD (EXP m n) = (ODD m | n = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   591
  by (import hollight ODD_EXP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   592
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   593
lemma EVEN_DOUBLE: "ALL n::nat. EVEN (NUMERAL_BIT0 (NUMERAL_BIT1 0) * n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   594
  by (import hollight EVEN_DOUBLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   595
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   596
lemma ODD_DOUBLE: "ALL x::nat. ODD (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * x))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   597
  by (import hollight ODD_DOUBLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   598
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   599
lemma EVEN_EXISTS_LEMMA: "ALL n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   600
   (EVEN n --> (EX m::nat. n = NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   601
   (~ EVEN n --> (EX m::nat. n = Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   602
  by (import hollight EVEN_EXISTS_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   603
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   604
lemma EVEN_EXISTS: "ALL n::nat. EVEN n = (EX m::nat. n = NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   605
  by (import hollight EVEN_EXISTS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   606
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   607
lemma ODD_EXISTS: "ALL n::nat. ODD n = (EX m::nat. n = Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * m))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   608
  by (import hollight ODD_EXISTS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   609
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   610
lemma EVEN_ODD_DECOMPOSITION: "ALL n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   611
   (EX (k::nat) m::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   612
       ODD m & n = EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) k * m) =
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   613
   (n ~= 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   614
  by (import hollight EVEN_ODD_DECOMPOSITION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   615
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   616
lemma SUB_0: "ALL x::nat. 0 - x = 0 & x - 0 = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   617
  by (import hollight SUB_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   618
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   619
lemma SUB_PRESUC: "ALL (m::nat) n::nat. Pred (Suc m - n) = m - n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   620
  by (import hollight SUB_PRESUC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   621
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   622
lemma SUB_EQ_0: "ALL (m::nat) n::nat. (m - n = 0) = <= m n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   623
  by (import hollight SUB_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   624
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   625
lemma ADD_SUBR: "ALL (x::nat) xa::nat. xa - (x + xa) = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   626
  by (import hollight ADD_SUBR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   627
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   628
lemma SUB_ADD: "ALL (x::nat) xa::nat. <= xa x --> x - xa + xa = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   629
  by (import hollight SUB_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   630
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   631
lemma SUC_SUB1: "ALL x::nat. Suc x - NUMERAL_BIT1 0 = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   632
  by (import hollight SUC_SUB1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   633
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   634
definition FACT :: "nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   635
  "FACT ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   636
SOME FACT::nat => nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   637
   FACT 0 = NUMERAL_BIT1 0 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   638
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   639
lemma DEF_FACT: "FACT =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   640
(SOME FACT::nat => nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   641
    FACT 0 = NUMERAL_BIT1 0 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   642
  by (import hollight DEF_FACT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   643
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   644
lemma FACT_LT: "ALL n::nat. < 0 (FACT n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   645
  by (import hollight FACT_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   646
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   647
lemma FACT_LE: "ALL x::nat. <= (NUMERAL_BIT1 0) (FACT x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   648
  by (import hollight FACT_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   649
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   650
lemma FACT_MONO: "ALL (m::nat) n::nat. <= m n --> <= (FACT m) (FACT n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   651
  by (import hollight FACT_MONO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   652
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   653
lemma DIVMOD_EXIST: "ALL (m::nat) n::nat. n ~= 0 --> (EX (q::nat) r::nat. m = q * n + r & < r n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   654
  by (import hollight DIVMOD_EXIST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   655
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   656
lemma DIVMOD_EXIST_0: "ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   657
   EX (x::nat) xa::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   658
      COND (n = 0) (x = 0 & xa = 0) (m = x * n + xa & < xa n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   659
  by (import hollight DIVMOD_EXIST_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   660
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   661
definition DIV :: "nat => nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   662
  "DIV ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   663
SOME q::nat => nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   664
   EX r::nat => nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   665
      ALL (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   666
         COND (n = 0) (q m n = 0 & r m n = 0)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   667
          (m = q m n * n + r m n & < (r m n) n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   668
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   669
lemma DEF_DIV: "DIV =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   670
(SOME q::nat => nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   671
    EX r::nat => nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   672
       ALL (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   673
          COND (n = 0) (q m n = 0 & r m n = 0)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   674
           (m = q m n * n + r m n & < (r m n) n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   675
  by (import hollight DEF_DIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   676
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   677
definition MOD :: "nat => nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   678
  "MOD ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   679
SOME r::nat => nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   680
   ALL (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   681
      COND (n = 0) (DIV m n = 0 & r m n = 0)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   682
       (m = DIV m n * n + r m n & < (r m n) n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   683
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   684
lemma DEF_MOD: "MOD =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   685
(SOME r::nat => nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   686
    ALL (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   687
       COND (n = 0) (DIV m n = 0 & r m n = 0)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   688
        (m = DIV m n * n + r m n & < (r m n) n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   689
  by (import hollight DEF_MOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   690
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   691
lemma DIVISION: "ALL (m::nat) n::nat. n ~= 0 --> m = DIV m n * n + MOD m n & < (MOD m n) n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   692
  by (import hollight DIVISION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   693
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   694
lemma DIVMOD_UNIQ_LEMMA: "ALL (m::nat) (n::nat) (q1::nat) (r1::nat) (q2::nat) r2::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   695
   (m = q1 * n + r1 & < r1 n) & m = q2 * n + r2 & < r2 n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   696
   q1 = q2 & r1 = r2"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   697
  by (import hollight DIVMOD_UNIQ_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   698
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   699
lemma DIVMOD_UNIQ: "ALL (m::nat) (n::nat) (q::nat) r::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   700
   m = q * n + r & < r n --> DIV m n = q & MOD m n = r"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   701
  by (import hollight DIVMOD_UNIQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   702
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   703
lemma MOD_UNIQ: "ALL (m::nat) (n::nat) (q::nat) r::nat. m = q * n + r & < r n --> MOD m n = r"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   704
  by (import hollight MOD_UNIQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   705
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   706
lemma DIV_UNIQ: "ALL (m::nat) (n::nat) (q::nat) r::nat. m = q * n + r & < r n --> DIV m n = q"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   707
  by (import hollight DIV_UNIQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   708
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   709
lemma MOD_MULT: "ALL (x::nat) xa::nat. x ~= 0 --> MOD (x * xa) x = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   710
  by (import hollight MOD_MULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   711
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   712
lemma DIV_MULT: "ALL (x::nat) xa::nat. x ~= 0 --> DIV (x * xa) x = xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   713
  by (import hollight DIV_MULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   714
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   715
lemma DIV_DIV: "ALL (m::nat) (n::nat) p::nat. n * p ~= 0 --> DIV (DIV m n) p = DIV m (n * p)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   716
  by (import hollight DIV_DIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   717
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   718
lemma MOD_LT: "ALL (m::nat) n::nat. < m n --> MOD m n = m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   719
  by (import hollight MOD_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   720
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   721
lemma MOD_EQ: "ALL (m::nat) (n::nat) (p::nat) q::nat. m = n + q * p --> MOD m p = MOD n p"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   722
  by (import hollight MOD_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   723
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   724
lemma DIV_MOD: "ALL (m::nat) (n::nat) p::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   725
   n * p ~= 0 --> MOD (DIV m n) p = DIV (MOD m (n * p)) n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   726
  by (import hollight DIV_MOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   727
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   728
lemma DIV_1: "ALL n::nat. DIV n (NUMERAL_BIT1 0) = n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   729
  by (import hollight DIV_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   730
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   731
lemma EXP_LT_0: "ALL (x::nat) xa::nat. < 0 (EXP xa x) = (xa ~= 0 | x = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   732
  by (import hollight EXP_LT_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   733
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   734
lemma DIV_LE: "ALL (m::nat) n::nat. n ~= 0 --> <= (DIV m n) m"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   735
  by (import hollight DIV_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   736
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   737
lemma DIV_MUL_LE: "ALL (m::nat) n::nat. <= (n * DIV m n) m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   738
  by (import hollight DIV_MUL_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   739
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   740
lemma DIV_0: "ALL n::nat. n ~= 0 --> DIV 0 n = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   741
  by (import hollight DIV_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   742
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   743
lemma MOD_0: "ALL n::nat. n ~= 0 --> MOD 0 n = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   744
  by (import hollight MOD_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   745
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   746
lemma DIV_LT: "ALL (m::nat) n::nat. < m n --> DIV m n = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   747
  by (import hollight DIV_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   748
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   749
lemma MOD_MOD: "ALL (m::nat) (n::nat) p::nat. n * p ~= 0 --> MOD (MOD m (n * p)) n = MOD m n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   750
  by (import hollight MOD_MOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   751
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   752
lemma MOD_MOD_REFL: "ALL (m::nat) n::nat. n ~= 0 --> MOD (MOD m n) n = MOD m n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   753
  by (import hollight MOD_MOD_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   754
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   755
lemma DIV_MULT2: "ALL (x::nat) (xa::nat) xb::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   756
   x * xb ~= 0 --> DIV (x * xa) (x * xb) = DIV xa xb"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   757
  by (import hollight DIV_MULT2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   758
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   759
lemma MOD_MULT2: "ALL (x::nat) (xa::nat) xb::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   760
   x * xb ~= 0 --> MOD (x * xa) (x * xb) = x * MOD xa xb"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   761
  by (import hollight MOD_MULT2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   762
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   763
lemma MOD_1: "ALL n::nat. MOD n (NUMERAL_BIT1 0) = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   764
  by (import hollight MOD_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   765
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   766
lemma MOD_EXISTS: "ALL (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   767
   (EX q::nat. m = n * q) = COND (n = 0) (m = 0) (MOD m n = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   768
  by (import hollight MOD_EXISTS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   769
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   770
lemma LT_EXP: "ALL (x::nat) (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   771
   < (EXP x m) (EXP x n) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   772
   (<= (NUMERAL_BIT0 (NUMERAL_BIT1 0)) x & < m n | x = 0 & m ~= 0 & n = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   773
  by (import hollight LT_EXP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   774
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   775
lemma LE_EXP: "ALL (x::nat) (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   776
   <= (EXP x m) (EXP x n) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   777
   COND (x = 0) (m = 0 --> n = 0) (x = NUMERAL_BIT1 0 | <= m n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   778
  by (import hollight LE_EXP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   779
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   780
lemma DIV_MONO: "ALL (m::nat) (n::nat) p::nat. p ~= 0 & <= m n --> <= (DIV m p) (DIV n p)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   781
  by (import hollight DIV_MONO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   782
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   783
lemma DIV_MONO_LT: "ALL (m::nat) (n::nat) p::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   784
   p ~= 0 & <= (m + p) n --> < (DIV m p) (DIV n p)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   785
  by (import hollight DIV_MONO_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   786
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   787
lemma LE_LDIV: "ALL (a::nat) (b::nat) n::nat. a ~= 0 & <= b (a * n) --> <= (DIV b a) n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   788
  by (import hollight LE_LDIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   789
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   790
lemma LE_RDIV_EQ: "ALL (a::nat) (b::nat) n::nat. a ~= 0 --> <= n (DIV b a) = <= (a * n) b"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   791
  by (import hollight LE_RDIV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   792
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   793
lemma LE_LDIV_EQ: "ALL (a::nat) (b::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   794
   a ~= 0 --> <= (DIV b a) n = < b (a * (n + NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   795
  by (import hollight LE_LDIV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   796
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   797
lemma DIV_EQ_0: "ALL (m::nat) n::nat. n ~= 0 --> (DIV m n = 0) = < m n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   798
  by (import hollight DIV_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   799
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   800
lemma MOD_EQ_0: "ALL (m::nat) n::nat. n ~= 0 --> (MOD m n = 0) = (EX q::nat. m = q * n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   801
  by (import hollight MOD_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   802
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   803
lemma EVEN_MOD: "ALL n::nat. EVEN n = (MOD n (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   804
  by (import hollight EVEN_MOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   805
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   806
lemma ODD_MOD: "ALL n::nat. ODD n = (MOD n (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   807
  by (import hollight ODD_MOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   808
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   809
lemma MOD_MULT_RMOD: "ALL (m::nat) (n::nat) p::nat. n ~= 0 --> MOD (m * MOD p n) n = MOD (m * p) n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   810
  by (import hollight MOD_MULT_RMOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   811
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   812
lemma MOD_MULT_LMOD: "ALL (x::nat) (xa::nat) xb::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   813
   xa ~= 0 --> MOD (MOD x xa * xb) xa = MOD (x * xb) xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   814
  by (import hollight MOD_MULT_LMOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   815
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   816
lemma MOD_MULT_MOD2: "ALL (x::nat) (xa::nat) xb::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   817
   xa ~= 0 --> MOD (MOD x xa * MOD xb xa) xa = MOD (x * xb) xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   818
  by (import hollight MOD_MULT_MOD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   819
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   820
lemma MOD_EXP_MOD: "ALL (m::nat) (n::nat) p::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   821
   n ~= 0 --> MOD (EXP (MOD m n) p) n = MOD (EXP m p) n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   822
  by (import hollight MOD_EXP_MOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   823
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   824
lemma MOD_MULT_ADD: "ALL (m::nat) (n::nat) p::nat. MOD (m * n + p) n = MOD p n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   825
  by (import hollight MOD_MULT_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   826
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   827
lemma MOD_ADD_MOD: "ALL (a::nat) (b::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   828
   n ~= 0 --> MOD (MOD a n + MOD b n) n = MOD (a + b) n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   829
  by (import hollight MOD_ADD_MOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   830
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   831
lemma DIV_ADD_MOD: "ALL (a::nat) (b::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   832
   n ~= 0 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   833
   (MOD (a + b) n = MOD a n + MOD b n) = (DIV (a + b) n = DIV a n + DIV b n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   834
  by (import hollight DIV_ADD_MOD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   835
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   836
lemma DIV_REFL: "ALL n::nat. n ~= 0 --> DIV n n = NUMERAL_BIT1 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   837
  by (import hollight DIV_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   838
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   839
lemma MOD_LE: "ALL (m::nat) n::nat. n ~= 0 --> <= (MOD m n) m"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   840
  by (import hollight MOD_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   841
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   842
lemma DIV_MONO2: "ALL (m::nat) (n::nat) p::nat. p ~= 0 & <= p m --> <= (DIV n m) (DIV n p)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   843
  by (import hollight DIV_MONO2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   844
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   845
lemma DIV_LE_EXCLUSION: "ALL (a::nat) (b::nat) (c::nat) d::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   846
   b ~= 0 & < (b * c) ((a + NUMERAL_BIT1 0) * d) --> <= (DIV c d) (DIV a b)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   847
  by (import hollight DIV_LE_EXCLUSION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   848
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   849
lemma DIV_EQ_EXCLUSION: "< ((b::nat) * (c::nat)) (((a::nat) + NUMERAL_BIT1 0) * (d::nat)) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   850
< (a * d) ((c + NUMERAL_BIT1 0) * b) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   851
DIV a b = DIV c d"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   852
  by (import hollight DIV_EQ_EXCLUSION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   853
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   854
lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   855
(ALL x::nat. (b = a + x --> P 0) & (a = b + x --> P x))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   856
  by (import hollight SUB_ELIM_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   857
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   858
lemma PRE_ELIM_THM: "(P::nat => bool) (Pred (n::nat)) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   859
(ALL m::nat. (n = 0 --> P 0) & (n = Suc m --> P m))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   860
  by (import hollight PRE_ELIM_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   861
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   862
lemma DIVMOD_ELIM_THM: "(P::nat => nat => bool) (DIV (m::nat) (n::nat)) (MOD m n) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   863
(n = 0 & P 0 0 |
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   864
 n ~= 0 & (ALL (q::nat) r::nat. m = q * n + r & < r n --> P q r))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   865
  by (import hollight DIVMOD_ELIM_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   866
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   867
definition eqeq :: "'q_9910 => 'q_9909 => ('q_9910 => 'q_9909 => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   868
  "eqeq ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   869
%(u::'q_9910::type) (ua::'q_9909::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   870
   ub::'q_9910::type => 'q_9909::type => bool. ub u ua"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   871
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   872
lemma DEF__equal__equal_: "eqeq =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   873
(%(u::'q_9910::type) (ua::'q_9909::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   874
    ub::'q_9910::type => 'q_9909::type => bool. ub u ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   875
  by (import hollight DEF__equal__equal_)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   876
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   877
definition mod_nat :: "nat => nat => nat => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   878
  "mod_nat ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   879
%(u::nat) (ua::nat) ub::nat. EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   880
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   881
lemma DEF_mod_nat: "mod_nat =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   882
(%(u::nat) (ua::nat) ub::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   883
    EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   884
  by (import hollight DEF_mod_nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   885
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   886
definition minimal :: "(nat => bool) => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   887
  "minimal == %u::nat => bool. SOME n::nat. u n & (ALL m::nat. < m n --> ~ u m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   888
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   889
lemma DEF_minimal: "minimal =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   890
(%u::nat => bool. SOME n::nat. u n & (ALL m::nat. < m n --> ~ u m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   891
  by (import hollight DEF_minimal)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   892
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   893
lemma MINIMAL: "ALL P::nat => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   894
   Ex P = (P (minimal P) & (ALL x::nat. < x (minimal P) --> ~ P x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   895
  by (import hollight MINIMAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   896
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
   897
definition WF :: "('A => 'A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   898
  "WF ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   899
%u::'A::type => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   900
   ALL P::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   901
      Ex P --> (EX x::'A::type. P x & (ALL y::'A::type. u y x --> ~ P y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   902
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   903
lemma DEF_WF: "WF =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   904
(%u::'A::type => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   905
    ALL P::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   906
       Ex P --> (EX x::'A::type. P x & (ALL y::'A::type. u y x --> ~ P y)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   907
  by (import hollight DEF_WF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   908
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   909
lemma WF_EQ: "WF (u_353::'A::type => 'A::type => bool) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   910
(ALL P::'A::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   911
    Ex P = (EX x::'A::type. P x & (ALL y::'A::type. u_353 y x --> ~ P y)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   912
  by (import hollight WF_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   913
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   914
lemma WF_IND: "WF (u_353::'A::type => 'A::type => bool) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   915
(ALL P::'A::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   916
    (ALL x::'A::type. (ALL y::'A::type. u_353 y x --> P y) --> P x) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   917
    All P)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   918
  by (import hollight WF_IND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   919
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   920
lemma WF_DCHAIN: "WF (u_353::'A::type => 'A::type => bool) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   921
(~ (EX s::nat => 'A::type. ALL n::nat. u_353 (s (Suc n)) (s n)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   922
  by (import hollight WF_DCHAIN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   923
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   924
lemma WF_UREC: "WF (u_353::'A::type => 'A::type => bool) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   925
(ALL H::('A::type => 'B::type) => 'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   926
    (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   927
        (ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   928
    (ALL (f::'A::type => 'B::type) g::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   929
        (ALL x::'A::type. f x = H f x) & (ALL x::'A::type. g x = H g x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   930
        f = g))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   931
  by (import hollight WF_UREC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   932
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   933
lemma WF_UREC_WF: "(ALL H::('A::type => bool) => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   934
    (ALL (f::'A::type => bool) (g::'A::type => bool) x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   935
        (ALL z::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   936
            (u_353::'A::type => 'A::type => bool) z x --> f z = g z) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   937
        H f x = H g x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   938
    (ALL (f::'A::type => bool) g::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   939
        (ALL x::'A::type. f x = H f x) & (ALL x::'A::type. g x = H g x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   940
        f = g)) -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   941
WF u_353"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   942
  by (import hollight WF_UREC_WF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   943
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   944
lemma WF_REC_INVARIANT: "WF (u_353::'A::type => 'A::type => bool) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   945
(ALL (H::('A::type => 'B::type) => 'A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   946
    S::'A::type => 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   947
    (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   948
        (ALL z::'A::type. u_353 z x --> f z = g z & S z (f z)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   949
        H f x = H g x & S x (H f x)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   950
    (EX f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   951
  by (import hollight WF_REC_INVARIANT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   952
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   953
lemma WF_REC: "WF (u_353::'A::type => 'A::type => bool) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   954
(ALL H::('A::type => 'B::type) => 'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   955
    (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   956
        (ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   957
    (EX f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   958
  by (import hollight WF_REC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   959
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   960
lemma WF_REC_WF: "(ALL H::('A::type => nat) => 'A::type => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   961
    (ALL (f::'A::type => nat) (g::'A::type => nat) x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   962
        (ALL z::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   963
            (u_353::'A::type => 'A::type => bool) z x --> f z = g z) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   964
        H f x = H g x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   965
    (EX f::'A::type => nat. ALL x::'A::type. f x = H f x)) -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   966
WF u_353"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   967
  by (import hollight WF_REC_WF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   968
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   969
lemma WF_EREC: "WF (u_353::'A::type => 'A::type => bool) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   970
(ALL H::('A::type => 'B::type) => 'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   971
    (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   972
        (ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   973
    (EX! f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   974
  by (import hollight WF_EREC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   975
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   976
lemma WF_SUBSET: "(ALL (x::'A::type) y::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   977
    (u_353::'A::type => 'A::type => bool) x y -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   978
    (u_472::'A::type => 'A::type => bool) x y) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   979
WF u_472 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   980
WF u_353"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   981
  by (import hollight WF_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   982
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   983
lemma WF_MEASURE_GEN: "ALL m::'A::type => 'B::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   984
   WF (u_353::'B::type => 'B::type => bool) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
   985
   WF (%(x::'A::type) x'::'A::type. u_353 (m x) (m x'))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   986
  by (import hollight WF_MEASURE_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   987
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   988
lemma WF_LEX_DEPENDENT: "ALL (R::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   989
   S::'A::type => 'B::type => 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   990
   WF R & (ALL x::'A::type. WF (S x)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   991
   WF (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   992
        (%f::'A::type * 'B::type => 'A::type * 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   993
            ALL (r1::'A::type) s1::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   994
               GEQ (f (r1, s1))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   995
                (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   996
                  (%f::'A::type * 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   997
                      ALL (r2::'A::type) s2::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   998
                         GEQ (f (r2, s2))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
   999
                          (R r1 r2 | r1 = r2 & S r1 s1 s2)))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1000
  by (import hollight WF_LEX_DEPENDENT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1001
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1002
lemma WF_LEX: "ALL (x::'A::type => 'A::type => bool) xa::'B::type => 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1003
   WF x & WF xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1004
   WF (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1005
        (%f::'A::type * 'B::type => 'A::type * 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1006
            ALL (r1::'A::type) s1::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1007
               GEQ (f (r1, s1))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1008
                (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1009
                  (%f::'A::type * 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1010
                      ALL (r2::'A::type) s2::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1011
                         GEQ (f (r2, s2)) (x r1 r2 | r1 = r2 & xa s1 s2)))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1012
  by (import hollight WF_LEX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1013
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1014
lemma WF_POINTWISE: "WF (u_353::'A::type => 'A::type => bool) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1015
WF (u_472::'B::type => 'B::type => bool) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1016
WF (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1017
     (%f::'A::type * 'B::type => 'A::type * 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1018
         ALL (x1::'A::type) y1::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1019
            GEQ (f (x1, y1))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1020
             (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1021
               (%f::'A::type * 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1022
                   ALL (x2::'A::type) y2::'B::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1023
                      GEQ (f (x2, y2)) (u_353 x1 x2 & u_472 y1 y2)))))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1024
  by (import hollight WF_POINTWISE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1025
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1026
lemma WF_num: "WF <"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1027
  by (import hollight WF_num)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1028
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1029
lemma WF_REC_num: "ALL H::(nat => 'A::type) => nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1030
   (ALL (f::nat => 'A::type) (g::nat => 'A::type) x::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1031
       (ALL z::nat. < z x --> f z = g z) --> H f x = H g x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1032
   (EX f::nat => 'A::type. ALL x::nat. f x = H f x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1033
  by (import hollight WF_REC_num)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1034
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1035
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1036
  measure :: "('q_11107 => nat) => 'q_11107 => 'q_11107 => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1037
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1038
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1039
  measure_def: "hollight.measure ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1040
%(u::'q_11107::type => nat) (x::'q_11107::type) y::'q_11107::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1041
   < (u x) (u y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1042
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1043
lemma DEF_measure: "hollight.measure =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1044
(%(u::'q_11107::type => nat) (x::'q_11107::type) y::'q_11107::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1045
    < (u x) (u y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1046
  by (import hollight DEF_measure)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1047
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1048
lemma WF_MEASURE: "ALL m::'A::type => nat. WF (hollight.measure m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1049
  by (import hollight WF_MEASURE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1050
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1051
lemma MEASURE_LE: "(ALL x::'q_11137::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1052
    hollight.measure (m::'q_11137::type => nat) x (a::'q_11137::type) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1053
    hollight.measure m x (b::'q_11137::type)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1054
<= (m a) (m b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1055
  by (import hollight MEASURE_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1056
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1057
lemma WF_REFL: "ALL x::'A::type. WF (u_353::'A::type => 'A::type => bool) --> ~ u_353 x x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1058
  by (import hollight WF_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1059
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1060
lemma WF_FALSE: "WF (%(x::'A::type) y::'A::type. False)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1061
  by (import hollight WF_FALSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1062
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1063
lemma WF_REC_TAIL: "ALL (P::'A::type => bool) (g::'A::type => 'A::type) h::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1064
   EX f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1065
      ALL x::'A::type. f x = COND (P x) (f (g x)) (h x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1066
  by (import hollight WF_REC_TAIL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1067
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1068
lemma WF_REC_TAIL_GENERAL: "ALL (P::('A::type => 'B::type) => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1069
   (G::('A::type => 'B::type) => 'A::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1070
   H::('A::type => 'B::type) => 'A::type => 'B::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1071
   WF (u_353::'A::type => 'A::type => bool) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1072
   (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1073
       (ALL z::'A::type. u_353 z x --> f z = g z) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1074
       P f x = P g x & G f x = G g x & H f x = H g x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1075
   (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1076
       (ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1077
   (ALL (f::'A::type => 'B::type) (x::'A::type) y::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1078
       P f x & u_353 y (G f x) --> u_353 y x) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1079
   (EX f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1080
       ALL x::'A::type. f x = COND (P f x) (f (G f x)) (H f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1081
  by (import hollight WF_REC_TAIL_GENERAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1082
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1083
lemma ARITH_ZERO: "(op &::bool => bool => bool) ((op =::nat => nat => bool) (0::nat) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1084
 ((op =::nat => nat => bool) ((NUMERAL_BIT0::nat => nat) (0::nat)) (0::nat))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1085
  by (import hollight ARITH_ZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1086
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1087
lemma ARITH_SUC: "(ALL x::nat. Suc x = Suc x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1088
Suc 0 = NUMERAL_BIT1 0 &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1089
(ALL x::nat. Suc (NUMERAL_BIT0 x) = NUMERAL_BIT1 x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1090
(ALL x::nat. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT0 (Suc x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1091
  by (import hollight ARITH_SUC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1092
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1093
lemma ARITH_PRE: "(ALL x::nat. Pred x = Pred x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1094
Pred 0 = 0 &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1095
(ALL x::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1096
    Pred (NUMERAL_BIT0 x) = COND (x = 0) 0 (NUMERAL_BIT1 (Pred x))) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1097
(ALL x::nat. Pred (NUMERAL_BIT1 x) = NUMERAL_BIT0 x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1098
  by (import hollight ARITH_PRE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1099
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1100
lemma ARITH_ADD: "(op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1101
 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1102
   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1103
       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1104
        (%xa::nat.
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1105
            (op =::nat => nat => bool) ((plus::nat => nat => nat) x xa)
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1106
             ((plus::nat => nat => nat) x xa))))
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1107
 ((op &::bool => bool => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1108
   ((op =::nat => nat => bool) ((plus::nat => nat => nat) (0::nat) (0::nat))
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1109
     (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1110
   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1111
     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1112
       (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1113
           (op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1114
            ((plus::nat => nat => nat) (0::nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1115
              ((NUMERAL_BIT0::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1116
            ((NUMERAL_BIT0::nat => nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1117
     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1118
       ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1119
         (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1120
             (op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1121
              ((plus::nat => nat => nat) (0::nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1122
                ((NUMERAL_BIT1::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1123
              ((NUMERAL_BIT1::nat => nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1124
       ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1125
         ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1126
           (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1127
               (op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1128
                ((plus::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1129
                  (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1130
                ((NUMERAL_BIT0::nat => nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1131
         ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1132
           ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1133
             (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1134
                 (op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1135
                  ((plus::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1136
                    (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1137
                  ((NUMERAL_BIT1::nat => nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1138
           ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1139
             ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1140
               (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1141
                   (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1142
                    (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1143
                        (op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1144
                         ((plus::nat => nat => nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1145
                           ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1146
                           ((NUMERAL_BIT0::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1147
                         ((NUMERAL_BIT0::nat => nat)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1148
                           ((plus::nat => nat => nat) x xa)))))
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1149
             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1150
               ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1151
                 (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1152
                     (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1153
                      (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1154
                          (op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1155
                           ((plus::nat => nat => nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1156
                             ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1157
                             ((NUMERAL_BIT1::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1158
                           ((NUMERAL_BIT1::nat => nat)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1159
                             ((plus::nat => nat => nat) x xa)))))
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1160
               ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1161
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1162
                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1163
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1164
                        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1165
                            (op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1166
                             ((plus::nat => nat => nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1167
                               ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1168
                               ((NUMERAL_BIT0::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1169
                             ((NUMERAL_BIT1::nat => nat)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1170
                               ((plus::nat => nat => nat) x xa)))))
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1171
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1172
                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1173
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1174
                        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1175
                            (op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1176
                             ((plus::nat => nat => nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1177
                               ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1178
                               ((NUMERAL_BIT1::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1179
                             ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1180
                               ((Suc::nat => nat)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1181
                                 ((plus::nat => nat => nat) x
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1182
                                   xa))))))))))))))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1183
  by (import hollight ARITH_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1184
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1185
lemma ARITH_MULT: "(op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1186
 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1187
   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1188
       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1189
        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1190
            (op =::nat => nat => bool) ((op *::nat => nat => nat) x xa)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1191
             ((op *::nat => nat => nat) x xa))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1192
 ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1193
   ((op =::nat => nat => bool) ((op *::nat => nat => nat) (0::nat) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1194
     (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1195
   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1196
     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1197
       (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1198
           (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1199
            ((op *::nat => nat => nat) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1200
              ((NUMERAL_BIT0::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1201
            (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1202
     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1203
       ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1204
         (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1205
             (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1206
              ((op *::nat => nat => nat) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1207
                ((NUMERAL_BIT1::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1208
              (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1209
       ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1210
         ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1211
           (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1212
               (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1213
                ((op *::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1214
                  (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1215
                (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1216
         ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1217
           ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1218
             (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1219
                 (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1220
                  ((op *::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1221
                    (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1222
                  (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1223
           ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1224
             ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1225
               (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1226
                   (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1227
                    (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1228
                        (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1229
                         ((op *::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1230
                           ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1231
                           ((NUMERAL_BIT0::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1232
                         ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1233
                           ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1234
                             ((op *::nat => nat => nat) x xa))))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1235
             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1236
               ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1237
                 (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1238
                     (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1239
                      (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1240
                          (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1241
                           ((op *::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1242
                             ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1243
                             ((NUMERAL_BIT1::nat => nat) xa))
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1244
                           ((plus::nat => nat => nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1245
                             ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1246
                             ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1247
                               ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1248
                                 ((op *::nat => nat => nat) x xa)))))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1249
               ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1250
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1251
                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1252
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1253
                        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1254
                            (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1255
                             ((op *::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1256
                               ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1257
                               ((NUMERAL_BIT0::nat => nat) xa))
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1258
                             ((plus::nat => nat => nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1259
                               ((NUMERAL_BIT0::nat => nat) xa)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1260
                               ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1261
                                 ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1262
                                   ((op *::nat => nat => nat) x xa)))))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1263
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1264
                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1265
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1266
                        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1267
                            (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1268
                             ((op *::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1269
                               ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1270
                               ((NUMERAL_BIT1::nat => nat) xa))
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1271
                             ((plus::nat => nat => nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1272
                               ((NUMERAL_BIT1::nat => nat) x)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  1273
                               ((plus::nat => nat => nat)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1274
                                 ((NUMERAL_BIT0::nat => nat) xa)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1275
                                 ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1276
                                   ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1277
                                     ((op *::nat => nat => nat) x
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1278
 xa))))))))))))))))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1279
  by (import hollight ARITH_MULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1280
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1281
lemma ARITH_EXP: "(ALL (x::nat) xa::nat. EXP x xa = EXP x xa) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1282
EXP 0 0 = NUMERAL_BIT1 0 &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1283
(ALL m::nat. EXP (NUMERAL_BIT0 m) 0 = NUMERAL_BIT1 0) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1284
(ALL m::nat. EXP (NUMERAL_BIT1 m) 0 = NUMERAL_BIT1 0) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1285
(ALL n::nat. EXP 0 (NUMERAL_BIT0 n) = EXP 0 n * EXP 0 n) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1286
(ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1287
    EXP (NUMERAL_BIT0 m) (NUMERAL_BIT0 n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1288
    EXP (NUMERAL_BIT0 m) n * EXP (NUMERAL_BIT0 m) n) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1289
(ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1290
    EXP (NUMERAL_BIT1 m) (NUMERAL_BIT0 n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1291
    EXP (NUMERAL_BIT1 m) n * EXP (NUMERAL_BIT1 m) n) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1292
(ALL n::nat. EXP 0 (NUMERAL_BIT1 n) = 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1293
(ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1294
    EXP (NUMERAL_BIT0 m) (NUMERAL_BIT1 n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1295
    NUMERAL_BIT0 m * (EXP (NUMERAL_BIT0 m) n * EXP (NUMERAL_BIT0 m) n)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1296
(ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1297
    EXP (NUMERAL_BIT1 m) (NUMERAL_BIT1 n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1298
    NUMERAL_BIT1 m * (EXP (NUMERAL_BIT1 m) n * EXP (NUMERAL_BIT1 m) n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1299
  by (import hollight ARITH_EXP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1300
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1301
lemma ARITH_EVEN: "(ALL x::nat. EVEN x = EVEN x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1302
EVEN 0 = True &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1303
(ALL x::nat. EVEN (NUMERAL_BIT0 x) = True) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1304
(ALL x::nat. EVEN (NUMERAL_BIT1 x) = False)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1305
  by (import hollight ARITH_EVEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1306
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1307
lemma ARITH_ODD: "(ALL x::nat. ODD x = ODD x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1308
ODD 0 = False &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1309
(ALL x::nat. ODD (NUMERAL_BIT0 x) = False) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1310
(ALL x::nat. ODD (NUMERAL_BIT1 x) = True)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1311
  by (import hollight ARITH_ODD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1312
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1313
lemma ARITH_LE: "(ALL (x::nat) xa::nat. <= x xa = <= x xa) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1314
<= 0 0 = True &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1315
(ALL x::nat. <= (NUMERAL_BIT0 x) 0 = (x = 0)) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1316
(ALL x::nat. <= (NUMERAL_BIT1 x) 0 = False) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1317
(ALL x::nat. <= 0 (NUMERAL_BIT0 x) = True) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1318
(ALL x::nat. <= 0 (NUMERAL_BIT1 x) = True) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1319
(ALL (x::nat) xa::nat. <= (NUMERAL_BIT0 x) (NUMERAL_BIT0 xa) = <= x xa) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1320
(ALL (x::nat) xa::nat. <= (NUMERAL_BIT0 x) (NUMERAL_BIT1 xa) = <= x xa) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1321
(ALL (x::nat) xa::nat. <= (NUMERAL_BIT1 x) (NUMERAL_BIT0 xa) = < x xa) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1322
(ALL (x::nat) xa::nat. <= (NUMERAL_BIT1 x) (NUMERAL_BIT1 xa) = <= x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1323
  by (import hollight ARITH_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1324
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1325
lemma ARITH_LT: "(ALL (x::nat) xa::nat. < x xa = < x xa) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1326
< 0 0 = False &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1327
(ALL x::nat. < (NUMERAL_BIT0 x) 0 = False) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1328
(ALL x::nat. < (NUMERAL_BIT1 x) 0 = False) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1329
(ALL x::nat. < 0 (NUMERAL_BIT0 x) = < 0 x) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1330
(ALL x::nat. < 0 (NUMERAL_BIT1 x) = True) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1331
(ALL (x::nat) xa::nat. < (NUMERAL_BIT0 x) (NUMERAL_BIT0 xa) = < x xa) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1332
(ALL (x::nat) xa::nat. < (NUMERAL_BIT0 x) (NUMERAL_BIT1 xa) = <= x xa) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1333
(ALL (x::nat) xa::nat. < (NUMERAL_BIT1 x) (NUMERAL_BIT0 xa) = < x xa) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1334
(ALL (x::nat) xa::nat. < (NUMERAL_BIT1 x) (NUMERAL_BIT1 xa) = < x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1335
  by (import hollight ARITH_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1336
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1337
lemma ARITH_EQ: "(op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1338
 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1339
   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1340
       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1341
        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1342
            (op =::bool => bool => bool) ((op =::nat => nat => bool) x xa)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1343
             ((op =::nat => nat => bool) x xa))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1344
 ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1345
   ((op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1346
     ((op =::nat => nat => bool) (0::nat) (0::nat)) (True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1347
   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1348
     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1349
       (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1350
           (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1351
            ((op =::nat => nat => bool) ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1352
              (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1353
            ((op =::nat => nat => bool) x (0::nat))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1354
     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1355
       ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1356
         (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1357
             (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1358
              ((op =::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1359
                (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1360
              (False::bool)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1361
       ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1362
         ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1363
           (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1364
               (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1365
                ((op =::nat => nat => bool) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1366
                  ((NUMERAL_BIT0::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1367
                ((op =::nat => nat => bool) (0::nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1368
         ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1369
           ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1370
             (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1371
                 (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1372
                  ((op =::nat => nat => bool) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1373
                    ((NUMERAL_BIT1::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1374
                  (False::bool)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1375
           ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1376
             ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1377
               (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1378
                   (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1379
                    (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1380
                        (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1381
                         ((op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1382
                           ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1383
                           ((NUMERAL_BIT0::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1384
                         ((op =::nat => nat => bool) x xa))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1385
             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1386
               ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1387
                 (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1388
                     (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1389
                      (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1390
                          (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1391
                           ((op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1392
                             ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1393
                             ((NUMERAL_BIT1::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1394
                           (False::bool))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1395
               ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1396
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1397
                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1398
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1399
                        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1400
                            (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1401
                             ((op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1402
                               ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1403
                               ((NUMERAL_BIT0::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1404
                             (False::bool))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1405
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1406
                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1407
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1408
                        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1409
                            (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1410
                             ((op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1411
                               ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1412
                               ((NUMERAL_BIT1::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1413
                             ((op =::nat => nat => bool) x xa))))))))))))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1414
  by (import hollight ARITH_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1415
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1416
lemma ARITH_SUB: "(op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1417
 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1418
   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1419
       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1420
        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1421
            (op =::nat => nat => bool) ((op -::nat => nat => nat) x xa)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1422
             ((op -::nat => nat => nat) x xa))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1423
 ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1424
   ((op =::nat => nat => bool) ((op -::nat => nat => nat) (0::nat) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1425
     (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1426
   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1427
     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1428
       (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1429
           (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1430
            ((op -::nat => nat => nat) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1431
              ((NUMERAL_BIT0::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1432
            (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1433
     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1434
       ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1435
         (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1436
             (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1437
              ((op -::nat => nat => nat) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1438
                ((NUMERAL_BIT1::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1439
              (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1440
       ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1441
         ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1442
           (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1443
               (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1444
                ((op -::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1445
                  (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1446
                ((NUMERAL_BIT0::nat => nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1447
         ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1448
           ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1449
             (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1450
                 (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1451
                  ((op -::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1452
                    (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1453
                  ((NUMERAL_BIT1::nat => nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1454
           ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1455
             ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1456
               (%m::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1457
                   (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1458
                    (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1459
                        (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1460
                         ((op -::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1461
                           ((NUMERAL_BIT0::nat => nat) m)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1462
                           ((NUMERAL_BIT0::nat => nat) n))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1463
                         ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1464
                           ((op -::nat => nat => nat) m n)))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1465
             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1466
               ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1467
                 (%m::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1468
                     (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1469
                      (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1470
                          (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1471
                           ((op -::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1472
                             ((NUMERAL_BIT0::nat => nat) m)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1473
                             ((NUMERAL_BIT1::nat => nat) n))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1474
                           ((Pred::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1475
                             ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1476
                               ((op -::nat => nat => nat) m n))))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1477
               ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1478
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1479
                   (%m::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1480
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1481
                        (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1482
                            (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1483
                             ((op -::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1484
                               ((NUMERAL_BIT1::nat => nat) m)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1485
                               ((NUMERAL_BIT0::nat => nat) n))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1486
                             ((COND::bool => nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1487
                               ((<=::nat => nat => bool) n m)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1488
                               ((NUMERAL_BIT1::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1489
                                 ((op -::nat => nat => nat) m n))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1490
                               (0::nat)))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1491
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1492
                   (%m::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1493
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1494
                        (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1495
                            (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1496
                             ((op -::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1497
                               ((NUMERAL_BIT1::nat => nat) m)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1498
                               ((NUMERAL_BIT1::nat => nat) n))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1499
                             ((NUMERAL_BIT0::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1500
                               ((op -::nat => nat => nat) m n)))))))))))))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1501
  by (import hollight ARITH_SUB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1502
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1503
lemma right_th: "(s::nat) * NUMERAL_BIT1 (x::nat) = s + NUMERAL_BIT0 (s * x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1504
  by (import hollight right_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1505
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1506
lemma SEMIRING_PTHS: "(ALL (x::'A::type) (y::'A::type) z::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1507
    (add::'A::type => 'A::type => 'A::type) x (add y z) = add (add x y) z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1508
(ALL (x::'A::type) y::'A::type. add x y = add y x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1509
(ALL x::'A::type. add (r0::'A::type) x = x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1510
(ALL (x::'A::type) (y::'A::type) z::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1511
    (mul::'A::type => 'A::type => 'A::type) x (mul y z) = mul (mul x y) z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1512
(ALL (x::'A::type) y::'A::type. mul x y = mul y x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1513
(ALL x::'A::type. mul (r1::'A::type) x = x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1514
(ALL x::'A::type. mul r0 x = r0) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1515
(ALL (x::'A::type) (y::'A::type) z::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1516
    mul x (add y z) = add (mul x y) (mul x z)) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1517
(ALL x::'A::type. (pwr::'A::type => nat => 'A::type) x 0 = r1) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1518
(ALL (x::'A::type) n::nat. pwr x (Suc n) = mul x (pwr x n)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1519
mul r1 (x::'A::type) = x &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1520
add (mul (a::'A::type) (m::'A::type)) (mul (b::'A::type) m) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1521
mul (add a b) m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1522
add (mul a m) m = mul (add a r1) m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1523
add m (mul a m) = mul (add a r1) m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1524
add m m = mul (add r1 r1) m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1525
mul r0 m = r0 &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1526
add r0 a = a &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1527
add a r0 = a &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1528
mul a b = mul b a &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1529
mul (add a b) (c::'A::type) = add (mul a c) (mul b c) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1530
mul r0 a = r0 &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1531
mul a r0 = r0 &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1532
mul r1 a = a &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1533
mul a r1 = a &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1534
mul (mul (lx::'A::type) (ly::'A::type))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1535
 (mul (rx::'A::type) (ry::'A::type)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1536
mul (mul lx rx) (mul ly ry) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1537
mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1538
mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1539
mul (mul lx ly) rx = mul (mul lx rx) ly &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1540
mul (mul lx ly) rx = mul lx (mul ly rx) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1541
mul lx rx = mul rx lx &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1542
mul lx (mul rx ry) = mul (mul lx rx) ry &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1543
mul lx (mul rx ry) = mul rx (mul lx ry) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1544
add (add a b) (add c (d::'A::type)) = add (add a c) (add b d) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1545
add (add a b) c = add a (add b c) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1546
add a (add c d) = add c (add a d) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1547
add (add a b) c = add (add a c) b &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1548
add a c = add c a &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1549
add a (add c d) = add (add a c) d &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1550
mul (pwr x (p::nat)) (pwr x (q::nat)) = pwr x (p + q) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1551
mul x (pwr x q) = pwr x (Suc q) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1552
mul (pwr x q) x = pwr x (Suc q) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1553
mul x x = pwr x (NUMERAL_BIT0 (NUMERAL_BIT1 0)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1554
pwr (mul x (y::'A::type)) q = mul (pwr x q) (pwr y q) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1555
pwr (pwr x p) q = pwr x (p * q) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1556
pwr x 0 = r1 &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1557
pwr x (NUMERAL_BIT1 0) = x &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1558
mul x (add y (z::'A::type)) = add (mul x y) (mul x z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1559
pwr x (Suc q) = mul x (pwr x q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1560
  by (import hollight SEMIRING_PTHS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1561
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1562
lemma sth: "(ALL (x::nat) (y::nat) z::nat. x + (y + z) = x + y + z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1563
(ALL (x::nat) y::nat. x + y = y + x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1564
(ALL x::nat. 0 + x = x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1565
(ALL (x::nat) (y::nat) z::nat. x * (y * z) = x * y * z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1566
(ALL (x::nat) y::nat. x * y = y * x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1567
(ALL x::nat. NUMERAL_BIT1 0 * x = x) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1568
(ALL x::nat. 0 * x = 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1569
(ALL (x::nat) (xa::nat) xb::nat. x * (xa + xb) = x * xa + x * xb) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1570
(ALL x::nat. EXP x 0 = NUMERAL_BIT1 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1571
(ALL (x::nat) xa::nat. EXP x (Suc xa) = x * EXP x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1572
  by (import hollight sth)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1573
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1574
lemma NUM_INTEGRAL_LEMMA: "(w::nat) = (x::nat) + (d::nat) & (y::nat) = (z::nat) + (e::nat) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1575
(w * y + x * z = w * z + x * y) = (w = x | y = z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1576
  by (import hollight NUM_INTEGRAL_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1577
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1578
lemma NUM_INTEGRAL: "(ALL x::nat. 0 * x = 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1579
(ALL (x::nat) (xa::nat) xb::nat. (x + xa = x + xb) = (xa = xb)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1580
(ALL (w::nat) (x::nat) (y::nat) z::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1581
    (w * y + x * z = w * z + x * y) = (w = x | y = z))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1582
  by (import hollight NUM_INTEGRAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1583
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1584
lemma INJ_INVERSE2: "ALL P::'A::type => 'B::type => 'C::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1585
   (ALL (x1::'A::type) (y1::'B::type) (x2::'A::type) y2::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1586
       (P x1 y1 = P x2 y2) = (x1 = x2 & y1 = y2)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1587
   (EX (x::'C::type => 'A::type) Y::'C::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1588
       ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1589
  by (import hollight INJ_INVERSE2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1590
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1591
definition NUMPAIR :: "nat => nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1592
  "NUMPAIR ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1593
%(u::nat) ua::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1594
   EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) u *
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1595
   (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua + NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1596
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1597
lemma DEF_NUMPAIR: "NUMPAIR =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1598
(%(u::nat) ua::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1599
    EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) u *
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1600
    (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua + NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1601
  by (import hollight DEF_NUMPAIR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1602
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1603
lemma NUMPAIR_INJ_LEMMA: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1604
   NUMPAIR x xa = NUMPAIR xb xc --> x = xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1605
  by (import hollight NUMPAIR_INJ_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1606
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1607
lemma NUMPAIR_INJ: "ALL (x1::nat) (y1::nat) (x2::nat) y2::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1608
   (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1609
  by (import hollight NUMPAIR_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1610
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1611
definition NUMFST :: "nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1612
  "NUMFST ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1613
SOME X::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1614
   EX Y::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1615
      ALL (x::nat) y::nat. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1616
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1617
lemma DEF_NUMFST: "NUMFST =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1618
(SOME X::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1619
    EX Y::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1620
       ALL (x::nat) y::nat. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1621
  by (import hollight DEF_NUMFST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1622
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1623
definition NUMSND :: "nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1624
  "NUMSND ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1625
SOME Y::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1626
   ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1627
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1628
lemma DEF_NUMSND: "NUMSND =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1629
(SOME Y::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1630
    ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1631
  by (import hollight DEF_NUMSND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1632
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1633
definition NUMSUM :: "bool => nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1634
  "NUMSUM ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1635
%(u::bool) ua::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1636
   COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1637
    (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1638
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1639
lemma DEF_NUMSUM: "NUMSUM =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1640
(%(u::bool) ua::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1641
    COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1642
     (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1643
  by (import hollight DEF_NUMSUM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1645
lemma NUMSUM_INJ: "ALL (b1::bool) (x1::nat) (b2::bool) x2::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1646
   (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1647
  by (import hollight NUMSUM_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1648
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1649
definition NUMLEFT :: "nat => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1650
  "NUMLEFT ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1651
SOME X::nat => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1652
   EX Y::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1653
      ALL (x::bool) y::nat. X (NUMSUM x y) = x & Y (NUMSUM x y) = y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1654
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1655
lemma DEF_NUMLEFT: "NUMLEFT =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1656
(SOME X::nat => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1657
    EX Y::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1658
       ALL (x::bool) y::nat. X (NUMSUM x y) = x & Y (NUMSUM x y) = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1659
  by (import hollight DEF_NUMLEFT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1660
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1661
definition NUMRIGHT :: "nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1662
  "NUMRIGHT ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1663
SOME Y::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1664
   ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1665
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1666
lemma DEF_NUMRIGHT: "NUMRIGHT =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1667
(SOME Y::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1668
    ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1669
  by (import hollight DEF_NUMRIGHT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1670
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1671
definition INJN :: "nat => nat => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1672
  "INJN == %(u::nat) (n::nat) a::'A::type. n = u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1673
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1674
lemma DEF_INJN: "INJN = (%(u::nat) (n::nat) a::'A::type. n = u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1675
  by (import hollight DEF_INJN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1676
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1677
lemma INJN_INJ: "(All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1678
 (%n1::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1679
     (All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1680
      (%n2::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1681
          (op =::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1682
           ((op =::(nat => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1683
                   => (nat => 'A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1684
             ((INJN::nat => nat => 'A::type => bool) n1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1685
             ((INJN::nat => nat => 'A::type => bool) n2))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1686
           ((op =::nat => nat => bool) n1 n2)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1687
  by (import hollight INJN_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1688
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1689
definition INJA :: "'A => nat => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1690
  "INJA == %(u::'A::type) (n::nat) b::'A::type. b = u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1691
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1692
lemma DEF_INJA: "INJA = (%(u::'A::type) (n::nat) b::'A::type. b = u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1693
  by (import hollight DEF_INJA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1694
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1695
lemma INJA_INJ: "ALL (a1::'A::type) a2::'A::type. (INJA a1 = INJA a2) = (a1 = a2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1696
  by (import hollight INJA_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1697
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1698
definition INJF :: "(nat => nat => 'A => bool) => nat => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1699
  "INJF == %(u::nat => nat => 'A::type => bool) n::nat. u (NUMFST n) (NUMSND n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1700
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1701
lemma DEF_INJF: "INJF =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1702
(%(u::nat => nat => 'A::type => bool) n::nat. u (NUMFST n) (NUMSND n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1703
  by (import hollight DEF_INJF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1704
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1705
lemma INJF_INJ: "ALL (f1::nat => nat => 'A::type => bool) f2::nat => nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1706
   (INJF f1 = INJF f2) = (f1 = f2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1707
  by (import hollight INJF_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1708
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1709
definition INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1710
  "INJP ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1711
%(u::nat => 'A::type => bool) (ua::nat => 'A::type => bool) (n::nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1712
   a::'A::type. COND (NUMLEFT n) (u (NUMRIGHT n) a) (ua (NUMRIGHT n) a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1713
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1714
lemma DEF_INJP: "INJP =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1715
(%(u::nat => 'A::type => bool) (ua::nat => 'A::type => bool) (n::nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1716
    a::'A::type. COND (NUMLEFT n) (u (NUMRIGHT n) a) (ua (NUMRIGHT n) a))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1717
  by (import hollight DEF_INJP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1718
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1719
lemma INJP_INJ: "ALL (f1::nat => 'A::type => bool) (f1'::nat => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1720
   (f2::nat => 'A::type => bool) f2'::nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1721
   (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1722
  by (import hollight INJP_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1723
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1724
definition ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1725
  "ZCONSTR ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1726
%(u::nat) (ua::'A::type) ub::nat => nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1727
   INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1728
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1729
lemma DEF_ZCONSTR: "ZCONSTR =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1730
(%(u::nat) (ua::'A::type) ub::nat => nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1731
    INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1732
  by (import hollight DEF_ZCONSTR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1733
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1734
definition ZBOT :: "nat => 'A => bool" where 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1735
  "ZBOT == INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1736
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1737
lemma DEF_ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1738
  by (import hollight DEF_ZBOT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1739
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1740
lemma ZCONSTR_ZBOT: "ALL (x::nat) (xa::'A::type) xb::nat => nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1741
   ZCONSTR x xa xb ~= ZBOT"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1742
  by (import hollight ZCONSTR_ZBOT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1743
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1744
definition ZRECSPACE :: "(nat => 'A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1745
  "ZRECSPACE ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1746
%a::nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1747
   ALL ZRECSPACE'::(nat => 'A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1748
      (ALL a::nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1749
          a = ZBOT |
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1750
          (EX (c::nat) (i::'A::type) r::nat => nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1751
              a = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1752
          ZRECSPACE' a) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1753
      ZRECSPACE' a"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1754
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1755
lemma DEF_ZRECSPACE: "ZRECSPACE =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1756
(%a::nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1757
    ALL ZRECSPACE'::(nat => 'A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1758
       (ALL a::nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1759
           a = ZBOT |
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1760
           (EX (c::nat) (i::'A::type) r::nat => nat => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1761
               a = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1762
           ZRECSPACE' a) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1763
       ZRECSPACE' a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1764
  by (import hollight DEF_ZRECSPACE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1765
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1766
typedef (open) ('A) recspace = "(Collect::((nat => 'A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1767
          => (nat => 'A::type => bool) set)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1768
 (ZRECSPACE::(nat => 'A::type => bool) => bool)"  morphisms "_dest_rec" "_mk_rec"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1769
  apply (rule light_ex_imp_nonempty[where t="ZBOT::nat => 'A::type => bool"])
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1770
  by (import hollight TYDEF_recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1771
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1772
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1773
  "_dest_rec" :: _ ("'_dest'_rec")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1774
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1775
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1776
  "_mk_rec" :: _ ("'_mk'_rec")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1777
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1778
lemmas "TYDEF_recspace_@intern" = typedef_hol2hollight 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1779
  [where a="a :: 'A recspace" and r=r ,
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1780
   OF type_definition_recspace]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1781
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1782
definition BOTTOM :: "'A recspace" where 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1783
  "(op ==::'A::type recspace => 'A::type recspace => prop)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1784
 (BOTTOM::'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1785
 ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1786
   (ZBOT::nat => 'A::type => bool))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1787
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1788
lemma DEF_BOTTOM: "(op =::'A::type recspace => 'A::type recspace => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1789
 (BOTTOM::'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1790
 ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1791
   (ZBOT::nat => 'A::type => bool))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1792
  by (import hollight DEF_BOTTOM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1793
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1794
definition CONSTR :: "nat => 'A => (nat => 'A recspace) => 'A recspace" where 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1795
  "(op ==::(nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1796
        => (nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1797
            => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1798
           => prop)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1799
 (CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1800
          => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1801
 (%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1802
     (_mk_rec::(nat => 'A::type => bool) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1803
      ((ZCONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1804
                 => 'A::type
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1805
                    => (nat => nat => 'A::type => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1806
                       => nat => 'A::type => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1807
        u ua
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1808
        (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1809
            (_dest_rec::'A::type recspace => nat => 'A::type => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1810
             (ub n))))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1811
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1812
lemma DEF_CONSTR: "(op =::(nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1813
       => (nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1814
           => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1815
          => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1816
 (CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1817
          => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1818
 (%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1819
     (_mk_rec::(nat => 'A::type => bool) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1820
      ((ZCONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1821
                 => 'A::type
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1822
                    => (nat => nat => 'A::type => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1823
                       => nat => 'A::type => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1824
        u ua
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1825
        (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1826
            (_dest_rec::'A::type recspace => nat => 'A::type => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1827
             (ub n))))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1828
  by (import hollight DEF_CONSTR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1829
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1830
lemma MK_REC_INJ: "(All::((nat => 'A::type => bool) => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1831
 (%x::nat => 'A::type => bool.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1832
     (All::((nat => 'A::type => bool) => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1833
      (%y::nat => 'A::type => bool.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1834
          (op -->::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1835
           ((op =::'A::type recspace => 'A::type recspace => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1836
             ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1837
             ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) y))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1838
           ((op -->::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1839
             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1840
               ((ZRECSPACE::(nat => 'A::type => bool) => bool) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1841
               ((ZRECSPACE::(nat => 'A::type => bool) => bool) y))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1842
             ((op =::(nat => 'A::type => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1843
                     => (nat => 'A::type => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1844
               x y))))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1845
  by (import hollight MK_REC_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1846
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1847
lemma CONSTR_BOT: "ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1848
   CONSTR c i r ~= BOTTOM"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1849
  by (import hollight CONSTR_BOT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1850
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1851
lemma CONSTR_INJ: "ALL (c1::nat) (i1::'A::type) (r1::nat => 'A::type recspace) (c2::nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1852
   (i2::'A::type) r2::nat => 'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1853
   (CONSTR c1 i1 r1 = CONSTR c2 i2 r2) = (c1 = c2 & i1 = i2 & r1 = r2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1854
  by (import hollight CONSTR_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1855
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1856
lemma CONSTR_IND: "ALL P::'A::type recspace => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1857
   P BOTTOM &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1858
   (ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1859
       (ALL n::nat. P (r n)) --> P (CONSTR c i r)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1860
   All P"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1861
  by (import hollight CONSTR_IND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1862
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1863
lemma CONSTR_REC: "ALL Fn::nat
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1864
        => 'A::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1865
           => (nat => 'A::type recspace) => (nat => 'B::type) => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1866
   EX f::'A::type recspace => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1867
      ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1868
         f (CONSTR c i r) = Fn c i r (%n::nat. f (r n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1869
  by (import hollight CONSTR_REC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1870
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1871
definition FCONS :: "'A => (nat => 'A) => nat => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1872
  "FCONS ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1873
SOME FCONS::'A::type => (nat => 'A::type) => nat => 'A::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1874
   (ALL (a::'A::type) f::nat => 'A::type. FCONS a f 0 = a) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1875
   (ALL (a::'A::type) (f::nat => 'A::type) n::nat. FCONS a f (Suc n) = f n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1876
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1877
lemma DEF_FCONS: "FCONS =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1878
(SOME FCONS::'A::type => (nat => 'A::type) => nat => 'A::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1879
    (ALL (a::'A::type) f::nat => 'A::type. FCONS a f 0 = a) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1880
    (ALL (a::'A::type) (f::nat => 'A::type) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1881
        FCONS a f (Suc n) = f n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1882
  by (import hollight DEF_FCONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1883
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1884
lemma FCONS_UNDO: "ALL f::nat => 'A::type. f = FCONS (f 0) (f o Suc)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1885
  by (import hollight FCONS_UNDO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1886
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1887
definition FNIL :: "nat => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1888
  "FNIL == %u::nat. SOME x::'A::type. True"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1889
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1890
lemma DEF_FNIL: "FNIL = (%u::nat. SOME x::'A::type. True)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1891
  by (import hollight DEF_FNIL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1892
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1893
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1894
  OUTL :: "'A + 'B => 'A" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1895
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1896
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1897
  OUTL_def: "hollight.OUTL ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1898
SOME OUTL::'A::type + 'B::type => 'A::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1899
   ALL x::'A::type. OUTL (Inl x) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1900
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1901
lemma DEF_OUTL: "hollight.OUTL =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1902
(SOME OUTL::'A::type + 'B::type => 'A::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1903
    ALL x::'A::type. OUTL (Inl x) = x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1904
  by (import hollight DEF_OUTL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1905
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1906
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1907
  OUTR :: "'A + 'B => 'B" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1908
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1909
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1910
  OUTR_def: "hollight.OUTR ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1911
SOME OUTR::'A::type + 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1912
   ALL y::'B::type. OUTR (Inr y) = y"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1913
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1914
lemma DEF_OUTR: "hollight.OUTR =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1915
(SOME OUTR::'A::type + 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  1916
    ALL y::'B::type. OUTR (Inr y) = y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1917
  by (import hollight DEF_OUTR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1918
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1919
typedef (open) ('A) option = "(Collect::('A::type recspace => bool) => 'A::type recspace set)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1920
 (%a::'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1921
     (All::(('A::type recspace => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1922
      (%option'::'A::type recspace => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1923
          (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1924
           ((All::('A::type recspace => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1925
             (%a::'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1926
                 (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1927
                  ((op |::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1928
                    ((op =::'A::type recspace => 'A::type recspace => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1929
                      a ((CONSTR::nat
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1930
                                  => 'A::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1931
                                     => (nat => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1932
  => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1933
                          ((NUMERAL::nat => nat) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1934
                          ((Eps::('A::type => bool) => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1935
                            (%v::'A::type. True::bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1936
                          (%n::nat. BOTTOM::'A::type recspace)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1937
                    ((Ex::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1938
                      (%aa::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1939
                          (op =::'A::type recspace
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1940
                                 => 'A::type recspace => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1941
                           a ((CONSTR::nat
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1942
 => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1943
                               ((Suc::nat => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1944
                                 ((NUMERAL::nat => nat) (0::nat)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1945
                               aa (%n::nat. BOTTOM::'A::type recspace)))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1946
                  (option' a)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1947
           (option' a)))"  morphisms "_dest_option" "_mk_option"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1948
  apply (rule light_ex_imp_nonempty[where t="(CONSTR::nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1949
 ((NUMERAL::nat => nat) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1950
 ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1951
 (%n::nat. BOTTOM::'A::type recspace)"])
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1952
  by (import hollight TYDEF_option)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1953
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1954
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1955
  "_dest_option" :: _ ("'_dest'_option")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1956
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1957
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1958
  "_mk_option" :: _ ("'_mk'_option")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1959
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1960
lemmas "TYDEF_option_@intern" = typedef_hol2hollight 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1961
  [where a="a :: 'A hollight.option" and r=r ,
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1962
   OF type_definition_option]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1963
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  1964
definition NONE :: "'A hollight.option" where 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1965
  "(op ==::'A::type hollight.option => 'A::type hollight.option => prop)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1966
 (NONE::'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1967
 ((_mk_option::'A::type recspace => 'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1968
   ((CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1969
             => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1970
     (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1971
     ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1972
     (%n::nat. BOTTOM::'A::type recspace)))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1973
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1974
lemma DEF_NONE: "(op =::'A::type hollight.option => 'A::type hollight.option => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1975
 (NONE::'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1976
 ((_mk_option::'A::type recspace => 'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1977
   ((CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1978
             => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1979
     (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1980
     ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1981
     (%n::nat. BOTTOM::'A::type recspace)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1982
  by (import hollight DEF_NONE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1983
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1984
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1985
  SOME :: "'A => 'A hollight.option" ("SOME")
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1986
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  1987
defs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1988
  SOME_def: "(op ==::('A::type => 'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1989
        => ('A::type => 'A::type hollight.option) => prop)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1990
 (SOME::'A::type => 'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1991
 (%a::'A::type.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1992
     (_mk_option::'A::type recspace => 'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1993
      ((CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1994
                => 'A::type
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1995
                   => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1996
        ((Suc::nat => nat) (0::nat)) a
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1997
        (%n::nat. BOTTOM::'A::type recspace)))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1998
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1999
lemma DEF_SOME: "(op =::('A::type => 'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2000
       => ('A::type => 'A::type hollight.option) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2001
 (SOME::'A::type => 'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2002
 (%a::'A::type.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2003
     (_mk_option::'A::type recspace => 'A::type hollight.option)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2004
      ((CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2005
                => 'A::type
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2006
                   => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2007
        ((Suc::nat => nat) (0::nat)) a
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2008
        (%n::nat. BOTTOM::'A::type recspace)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2009
  by (import hollight DEF_SOME)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2010
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2011
typedef (open) ('A) list = "(Collect::('A::type recspace => bool) => 'A::type recspace set)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2012
 (%a::'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2013
     (All::(('A::type recspace => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2014
      (%list'::'A::type recspace => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2015
          (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2016
           ((All::('A::type recspace => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2017
             (%a::'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2018
                 (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2019
                  ((op |::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2020
                    ((op =::'A::type recspace => 'A::type recspace => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2021
                      a ((CONSTR::nat
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2022
                                  => 'A::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2023
                                     => (nat => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2024
  => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2025
                          ((NUMERAL::nat => nat) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2026
                          ((Eps::('A::type => bool) => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2027
                            (%v::'A::type. True::bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2028
                          (%n::nat. BOTTOM::'A::type recspace)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2029
                    ((Ex::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2030
                      (%a0::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2031
                          (Ex::('A::type recspace => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2032
                           (%a1::'A::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2033
                               (op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2034
                                ((op =::'A::type recspace
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2035
  => 'A::type recspace => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2036
                                  a ((CONSTR::nat
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2037
        => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2038
((Suc::nat => nat) ((NUMERAL::nat => nat) (0::nat))) a0
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2039
((FCONS::'A::type recspace
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2040
         => (nat => 'A::type recspace) => nat => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2041
  a1 (%n::nat. BOTTOM::'A::type recspace))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2042
                                (list' a1)))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2043
                  (list' a)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2044
           (list' a)))"  morphisms "_dest_list" "_mk_list"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2045
  apply (rule light_ex_imp_nonempty[where t="(CONSTR::nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2046
 ((NUMERAL::nat => nat) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2047
 ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2048
 (%n::nat. BOTTOM::'A::type recspace)"])
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2049
  by (import hollight TYDEF_list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2050
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2051
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2052
  "_dest_list" :: _ ("'_dest'_list")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2053
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2054
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2055
  "_mk_list" :: _ ("'_mk'_list")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2056
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2057
lemmas "TYDEF_list_@intern" = typedef_hol2hollight 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2058
  [where a="a :: 'A hollight.list" and r=r ,
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2059
   OF type_definition_list]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2060
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2061
definition NIL :: "'A hollight.list" where 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2062
  "(op ==::'A::type hollight.list => 'A::type hollight.list => prop)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2063
 (NIL::'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2064
 ((_mk_list::'A::type recspace => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2065
   ((CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2066
             => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2067
     (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2068
     ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2069
     (%n::nat. BOTTOM::'A::type recspace)))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2070
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2071
lemma DEF_NIL: "(op =::'A::type hollight.list => 'A::type hollight.list => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2072
 (NIL::'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2073
 ((_mk_list::'A::type recspace => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2074
   ((CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2075
             => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2076
     (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2077
     ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2078
     (%n::nat. BOTTOM::'A::type recspace)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2079
  by (import hollight DEF_NIL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2080
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2081
definition CONS :: "'A => 'A hollight.list => 'A hollight.list" where 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2082
  "(op ==::('A::type => 'A::type hollight.list => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2083
        => ('A::type => 'A::type hollight.list => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2084
           => prop)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2085
 (CONS::'A::type => 'A::type hollight.list => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2086
 (%(a0::'A::type) a1::'A::type hollight.list.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2087
     (_mk_list::'A::type recspace => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2088
      ((CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2089
                => 'A::type
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2090
                   => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2091
        ((Suc::nat => nat) (0::nat)) a0
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2092
        ((FCONS::'A::type recspace
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2093
                 => (nat => 'A::type recspace) => nat => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2094
          ((_dest_list::'A::type hollight.list => 'A::type recspace) a1)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2095
          (%n::nat. BOTTOM::'A::type recspace))))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2096
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2097
lemma DEF_CONS: "(op =::('A::type => 'A::type hollight.list => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2098
       => ('A::type => 'A::type hollight.list => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2099
          => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2100
 (CONS::'A::type => 'A::type hollight.list => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2101
 (%(a0::'A::type) a1::'A::type hollight.list.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2102
     (_mk_list::'A::type recspace => 'A::type hollight.list)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2103
      ((CONSTR::nat
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2104
                => 'A::type
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2105
                   => (nat => 'A::type recspace) => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2106
        ((Suc::nat => nat) (0::nat)) a0
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2107
        ((FCONS::'A::type recspace
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2108
                 => (nat => 'A::type recspace) => nat => 'A::type recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2109
          ((_dest_list::'A::type hollight.list => 'A::type recspace) a1)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2110
          (%n::nat. BOTTOM::'A::type recspace))))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2111
  by (import hollight DEF_CONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2112
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2113
lemma pair_RECURSION: "ALL PAIR'::'A::type => 'B::type => 'C::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2114
   EX x::'A::type * 'B::type => 'C::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2115
      ALL (a0::'A::type) a1::'B::type. x (a0, a1) = PAIR' a0 a1"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2116
  by (import hollight pair_RECURSION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2117
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2118
lemma num_RECURSION_STD: "ALL (e::'Z::type) f::nat => 'Z::type => 'Z::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2119
   EX fn::nat => 'Z::type. fn 0 = e & (ALL n::nat. fn (Suc n) = f n (fn n))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2120
  by (import hollight num_RECURSION_STD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2121
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2122
lemma bool_RECURSION: "ALL (a::'A::type) b::'A::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2123
   EX x::bool => 'A::type. x False = a & x True = b"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2124
  by (import hollight bool_RECURSION)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2125
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2126
definition ISO :: "('A => 'B) => ('B => 'A) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2127
  "ISO ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2128
%(u::'A::type => 'B::type) ua::'B::type => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2129
   (ALL x::'B::type. u (ua x) = x) & (ALL y::'A::type. ua (u y) = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2130
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2131
lemma DEF_ISO: "ISO =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2132
(%(u::'A::type => 'B::type) ua::'B::type => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2133
    (ALL x::'B::type. u (ua x) = x) & (ALL y::'A::type. ua (u y) = y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2134
  by (import hollight DEF_ISO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2135
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2136
lemma ISO_REFL: "ISO (%x::'A::type. x) (%x::'A::type. x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2137
  by (import hollight ISO_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2138
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2139
lemma ISO_FUN: "ISO (f::'A::type => 'A'::type) (f'::'A'::type => 'A::type) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2140
ISO (g::'B::type => 'B'::type) (g'::'B'::type => 'B::type) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2141
ISO (%(h::'A::type => 'B::type) a'::'A'::type. g (h (f' a')))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2142
 (%(h::'A'::type => 'B'::type) a::'A::type. g' (h (f a)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2143
  by (import hollight ISO_FUN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2144
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2145
lemma ISO_USAGE: "ISO (f::'q_16636::type => 'q_16633::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2146
 (g::'q_16633::type => 'q_16636::type) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2147
(ALL P::'q_16636::type => bool. All P = (ALL x::'q_16633::type. P (g x))) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2148
(ALL P::'q_16636::type => bool. Ex P = (EX x::'q_16633::type. P (g x))) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2149
(ALL (a::'q_16636::type) b::'q_16633::type. (a = g b) = (f a = b))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2150
  by (import hollight ISO_USAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2151
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2152
typedef (open) N_2 = "{a::bool recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2153
 ALL u::bool recspace => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2154
    (ALL a::bool recspace.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2155
        a = CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM) |
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2156
        a =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2157
        CONSTR (Suc (NUMERAL 0)) (SOME x::bool. True) (%n::nat. BOTTOM) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2158
        u a) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2159
    u a}"  morphisms "_dest_2" "_mk_2"
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2160
  apply (rule light_ex_imp_nonempty[where t="CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM)"])
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2161
  by (import hollight TYDEF_2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2162
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2163
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2164
  "_dest_2" :: _ ("'_dest'_2")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2165
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2166
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2167
  "_mk_2" :: _ ("'_mk'_2")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2168
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2169
lemmas "TYDEF_2_@intern" = typedef_hol2hollight 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2170
  [where a="a :: N_2" and r=r ,
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2171
   OF type_definition_N_2]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2172
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2173
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2174
  "_10302" :: "N_2" ("'_10302")
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2175
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2176
defs
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2177
  "_10302_def": "(op ==::N_2 => N_2 => prop) (_10302::N_2)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2178
 ((_mk_2::bool recspace => N_2)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2179
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2180
     (0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2181
     (%n::nat. BOTTOM::bool recspace)))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2182
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2183
lemma DEF__10302: "(op =::N_2 => N_2 => bool) (_10302::N_2)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2184
 ((_mk_2::bool recspace => N_2)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2185
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2186
     (0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2187
     (%n::nat. BOTTOM::bool recspace)))"
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2188
  by (import hollight DEF__10302)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2189
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2190
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2191
  "_10303" :: "N_2" ("'_10303")
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2192
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2193
defs
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2194
  "_10303_def": "(op ==::N_2 => N_2 => prop) (_10303::N_2)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2195
 ((_mk_2::bool recspace => N_2)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2196
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2197
     ((Suc::nat => nat) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2198
     ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2199
     (%n::nat. BOTTOM::bool recspace)))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2200
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2201
lemma DEF__10303: "(op =::N_2 => N_2 => bool) (_10303::N_2)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2202
 ((_mk_2::bool recspace => N_2)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2203
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2204
     ((Suc::nat => nat) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2205
     ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2206
     (%n::nat. BOTTOM::bool recspace)))"
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2207
  by (import hollight DEF__10303)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2208
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2209
definition two_1 :: "N_2" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2210
  "two_1 == _10302"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2211
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2212
lemma DEF_two_1: "two_1 = _10302"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2213
  by (import hollight DEF_two_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2214
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2215
definition two_2 :: "N_2" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2216
  "two_2 == _10303"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2217
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2218
lemma DEF_two_2: "two_2 = _10303"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2219
  by (import hollight DEF_two_2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2220
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2221
typedef (open) N_3 = "{a::bool recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2222
 ALL u::bool recspace => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2223
    (ALL a::bool recspace.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2224
        a = CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM) |
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2225
        a =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2226
        CONSTR (Suc (NUMERAL 0)) (SOME x::bool. True) (%n::nat. BOTTOM) |
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2227
        a =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2228
        CONSTR (Suc (Suc (NUMERAL 0))) (SOME x::bool. True)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2229
         (%n::nat. BOTTOM) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2230
        u a) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2231
    u a}"  morphisms "_dest_3" "_mk_3"
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2232
  apply (rule light_ex_imp_nonempty[where t="CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM)"])
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2233
  by (import hollight TYDEF_3)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2234
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2235
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2236
  "_dest_3" :: _ ("'_dest'_3")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2237
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2238
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2239
  "_mk_3" :: _ ("'_mk'_3")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2240
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2241
lemmas "TYDEF_3_@intern" = typedef_hol2hollight 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2242
  [where a="a :: N_3" and r=r ,
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2243
   OF type_definition_N_3]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2244
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2245
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2246
  "_10326" :: "N_3" ("'_10326")
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2247
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2248
defs
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2249
  "_10326_def": "(op ==::N_3 => N_3 => prop) (_10326::N_3)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2250
 ((_mk_3::bool recspace => N_3)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2251
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2252
     (0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2253
     (%n::nat. BOTTOM::bool recspace)))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2254
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2255
lemma DEF__10326: "(op =::N_3 => N_3 => bool) (_10326::N_3)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2256
 ((_mk_3::bool recspace => N_3)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2257
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2258
     (0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2259
     (%n::nat. BOTTOM::bool recspace)))"
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2260
  by (import hollight DEF__10326)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2261
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2262
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2263
  "_10327" :: "N_3" ("'_10327")
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2264
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2265
defs
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2266
  "_10327_def": "(op ==::N_3 => N_3 => prop) (_10327::N_3)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2267
 ((_mk_3::bool recspace => N_3)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2268
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2269
     ((Suc::nat => nat) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2270
     ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2271
     (%n::nat. BOTTOM::bool recspace)))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2272
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2273
lemma DEF__10327: "(op =::N_3 => N_3 => bool) (_10327::N_3)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2274
 ((_mk_3::bool recspace => N_3)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2275
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2276
     ((Suc::nat => nat) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2277
     ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2278
     (%n::nat. BOTTOM::bool recspace)))"
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2279
  by (import hollight DEF__10327)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2280
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2281
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2282
  "_10328" :: "N_3" ("'_10328")
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2283
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2284
defs
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2285
  "_10328_def": "(op ==::N_3 => N_3 => prop) (_10328::N_3)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2286
 ((_mk_3::bool recspace => N_3)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2287
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2288
     ((Suc::nat => nat) ((Suc::nat => nat) (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2289
     ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2290
     (%n::nat. BOTTOM::bool recspace)))"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2291
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2292
lemma DEF__10328: "(op =::N_3 => N_3 => bool) (_10328::N_3)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2293
 ((_mk_3::bool recspace => N_3)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2294
   ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2295
     ((Suc::nat => nat) ((Suc::nat => nat) (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2296
     ((Eps::(bool => bool) => bool) (%x::bool. True::bool))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2297
     (%n::nat. BOTTOM::bool recspace)))"
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2298
  by (import hollight DEF__10328)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2299
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2300
definition three_1 :: "N_3" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2301
  "three_1 == _10326"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2302
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2303
lemma DEF_three_1: "three_1 = _10326"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2304
  by (import hollight DEF_three_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2305
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2306
definition three_2 :: "N_3" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2307
  "three_2 == _10327"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2308
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2309
lemma DEF_three_2: "three_2 = _10327"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2310
  by (import hollight DEF_three_2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2311
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2312
definition three_3 :: "N_3" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2313
  "three_3 == _10328"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2314
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2315
lemma DEF_three_3: "three_3 = _10328"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2316
  by (import hollight DEF_three_3)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2317
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2318
lemma list_INDUCT: "ALL P::'A::type hollight.list => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2319
   P NIL &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2320
   (ALL (a0::'A::type) a1::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2321
       P a1 --> P (CONS a0 a1)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2322
   All P"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2323
  by (import hollight list_INDUCT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2324
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2325
definition HD :: "'A hollight.list => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2326
  "HD ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2327
SOME HD::'A::type hollight.list => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2328
   ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2329
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2330
lemma DEF_HD: "HD =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2331
(SOME HD::'A::type hollight.list => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2332
    ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2333
  by (import hollight DEF_HD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2334
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2335
definition TL :: "'A hollight.list => 'A hollight.list" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2336
  "TL ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2337
SOME TL::'A::type hollight.list => 'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2338
   ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2339
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2340
lemma DEF_TL: "TL =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2341
(SOME TL::'A::type hollight.list => 'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2342
    ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2343
  by (import hollight DEF_TL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2344
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2345
definition APPEND :: "'A hollight.list => 'A hollight.list => 'A hollight.list" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2346
  "APPEND ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2347
SOME APPEND::'A::type hollight.list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2348
             => 'A::type hollight.list => 'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2349
   (ALL l::'A::type hollight.list. APPEND NIL l = l) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2350
   (ALL (h::'A::type) (t::'A::type hollight.list) l::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2351
       APPEND (CONS h t) l = CONS h (APPEND t l))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2352
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2353
lemma DEF_APPEND: "APPEND =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2354
(SOME APPEND::'A::type hollight.list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2355
              => 'A::type hollight.list => 'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2356
    (ALL l::'A::type hollight.list. APPEND NIL l = l) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2357
    (ALL (h::'A::type) (t::'A::type hollight.list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2358
        l::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2359
        APPEND (CONS h t) l = CONS h (APPEND t l)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2360
  by (import hollight DEF_APPEND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2361
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2362
definition REVERSE :: "'A hollight.list => 'A hollight.list" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2363
  "REVERSE ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2364
SOME REVERSE::'A::type hollight.list => 'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2365
   REVERSE NIL = NIL &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2366
   (ALL (l::'A::type hollight.list) x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2367
       REVERSE (CONS x l) = APPEND (REVERSE l) (CONS x NIL))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2368
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2369
lemma DEF_REVERSE: "REVERSE =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2370
(SOME REVERSE::'A::type hollight.list => 'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2371
    REVERSE NIL = NIL &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2372
    (ALL (l::'A::type hollight.list) x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2373
        REVERSE (CONS x l) = APPEND (REVERSE l) (CONS x NIL)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2374
  by (import hollight DEF_REVERSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2375
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2376
definition LENGTH :: "'A hollight.list => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2377
  "LENGTH ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2378
SOME LENGTH::'A::type hollight.list => nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2379
   LENGTH NIL = 0 &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2380
   (ALL (h::'A::type) t::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2381
       LENGTH (CONS h t) = Suc (LENGTH t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2382
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2383
lemma DEF_LENGTH: "LENGTH =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2384
(SOME LENGTH::'A::type hollight.list => nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2385
    LENGTH NIL = 0 &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2386
    (ALL (h::'A::type) t::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2387
        LENGTH (CONS h t) = Suc (LENGTH t)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2388
  by (import hollight DEF_LENGTH)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2389
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2390
definition MAP :: "('A => 'B) => 'A hollight.list => 'B hollight.list" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2391
  "MAP ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2392
SOME MAP::('A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2393
          => 'A::type hollight.list => 'B::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2394
   (ALL f::'A::type => 'B::type. MAP f NIL = NIL) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2395
   (ALL (f::'A::type => 'B::type) (h::'A::type) t::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2396
       MAP f (CONS h t) = CONS (f h) (MAP f t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2397
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2398
lemma DEF_MAP: "MAP =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2399
(SOME MAP::('A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2400
           => 'A::type hollight.list => 'B::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2401
    (ALL f::'A::type => 'B::type. MAP f NIL = NIL) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2402
    (ALL (f::'A::type => 'B::type) (h::'A::type) t::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2403
        MAP f (CONS h t) = CONS (f h) (MAP f t)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2404
  by (import hollight DEF_MAP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2405
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2406
definition LAST :: "'A hollight.list => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2407
  "LAST ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2408
SOME LAST::'A::type hollight.list => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2409
   ALL (h::'A::type) t::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2410
      LAST (CONS h t) = COND (t = NIL) h (LAST t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2411
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2412
lemma DEF_LAST: "LAST =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2413
(SOME LAST::'A::type hollight.list => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2414
    ALL (h::'A::type) t::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2415
       LAST (CONS h t) = COND (t = NIL) h (LAST t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2416
  by (import hollight DEF_LAST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2417
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2418
definition REPLICATE :: "nat => 'q_16860 => 'q_16860 hollight.list" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2419
  "REPLICATE ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2420
SOME REPLICATE::nat => 'q_16860::type => 'q_16860::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2421
   (ALL x::'q_16860::type. REPLICATE 0 x = NIL) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2422
   (ALL (n::nat) x::'q_16860::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2423
       REPLICATE (Suc n) x = CONS x (REPLICATE n x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2424
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2425
lemma DEF_REPLICATE: "REPLICATE =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2426
(SOME REPLICATE::nat => 'q_16860::type => 'q_16860::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2427
    (ALL x::'q_16860::type. REPLICATE 0 x = NIL) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2428
    (ALL (n::nat) x::'q_16860::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2429
        REPLICATE (Suc n) x = CONS x (REPLICATE n x)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2430
  by (import hollight DEF_REPLICATE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2431
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2432
definition NULL :: "'q_16875 hollight.list => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2433
  "NULL ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2434
SOME NULL::'q_16875::type hollight.list => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2435
   NULL NIL = True &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2436
   (ALL (h::'q_16875::type) t::'q_16875::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2437
       NULL (CONS h t) = False)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2438
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2439
lemma DEF_NULL: "NULL =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2440
(SOME NULL::'q_16875::type hollight.list => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2441
    NULL NIL = True &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2442
    (ALL (h::'q_16875::type) t::'q_16875::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2443
        NULL (CONS h t) = False))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2444
  by (import hollight DEF_NULL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2445
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2446
definition ALL_list :: "('q_16895 => bool) => 'q_16895 hollight.list => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2447
  "ALL_list ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2448
SOME u::('q_16895::type => bool) => 'q_16895::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2449
   (ALL P::'q_16895::type => bool. u P NIL = True) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2450
   (ALL (h::'q_16895::type) (P::'q_16895::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2451
       t::'q_16895::type hollight.list. u P (CONS h t) = (P h & u P t))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2452
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2453
lemma DEF_ALL: "ALL_list =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2454
(SOME u::('q_16895::type => bool) => 'q_16895::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2455
    (ALL P::'q_16895::type => bool. u P NIL = True) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2456
    (ALL (h::'q_16895::type) (P::'q_16895::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2457
        t::'q_16895::type hollight.list. u P (CONS h t) = (P h & u P t)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2458
  by (import hollight DEF_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2459
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2460
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2461
  EX :: "('q_16916 => bool) => 'q_16916 hollight.list => bool" ("EX")
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2462
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2463
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2464
  EX_def: "EX ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2465
SOME u::('q_16916::type => bool) => 'q_16916::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2466
   (ALL P::'q_16916::type => bool. u P NIL = False) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2467
   (ALL (h::'q_16916::type) (P::'q_16916::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2468
       t::'q_16916::type hollight.list. u P (CONS h t) = (P h | u P t))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2469
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2470
lemma DEF_EX: "EX =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2471
(SOME u::('q_16916::type => bool) => 'q_16916::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2472
    (ALL P::'q_16916::type => bool. u P NIL = False) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2473
    (ALL (h::'q_16916::type) (P::'q_16916::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2474
        t::'q_16916::type hollight.list. u P (CONS h t) = (P h | u P t)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2475
  by (import hollight DEF_EX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2476
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2477
definition ITLIST :: "('q_16939 => 'q_16938 => 'q_16938)
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2478
=> 'q_16939 hollight.list => 'q_16938 => 'q_16938" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2479
  "ITLIST ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2480
SOME ITLIST::('q_16939::type => 'q_16938::type => 'q_16938::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2481
             => 'q_16939::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2482
                => 'q_16938::type => 'q_16938::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2483
   (ALL (f::'q_16939::type => 'q_16938::type => 'q_16938::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2484
       b::'q_16938::type. ITLIST f NIL b = b) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2485
   (ALL (h::'q_16939::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2486
       (f::'q_16939::type => 'q_16938::type => 'q_16938::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2487
       (t::'q_16939::type hollight.list) b::'q_16938::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2488
       ITLIST f (CONS h t) b = f h (ITLIST f t b))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2489
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2490
lemma DEF_ITLIST: "ITLIST =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2491
(SOME ITLIST::('q_16939::type => 'q_16938::type => 'q_16938::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2492
              => 'q_16939::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2493
                 => 'q_16938::type => 'q_16938::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2494
    (ALL (f::'q_16939::type => 'q_16938::type => 'q_16938::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2495
        b::'q_16938::type. ITLIST f NIL b = b) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2496
    (ALL (h::'q_16939::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2497
        (f::'q_16939::type => 'q_16938::type => 'q_16938::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2498
        (t::'q_16939::type hollight.list) b::'q_16938::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2499
        ITLIST f (CONS h t) b = f h (ITLIST f t b)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2500
  by (import hollight DEF_ITLIST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2501
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2502
definition MEM :: "'q_16964 => 'q_16964 hollight.list => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2503
  "MEM ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2504
SOME MEM::'q_16964::type => 'q_16964::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2505
   (ALL x::'q_16964::type. MEM x NIL = False) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2506
   (ALL (h::'q_16964::type) (x::'q_16964::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2507
       t::'q_16964::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2508
       MEM x (CONS h t) = (x = h | MEM x t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2509
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2510
lemma DEF_MEM: "MEM =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2511
(SOME MEM::'q_16964::type => 'q_16964::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2512
    (ALL x::'q_16964::type. MEM x NIL = False) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2513
    (ALL (h::'q_16964::type) (x::'q_16964::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2514
        t::'q_16964::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2515
        MEM x (CONS h t) = (x = h | MEM x t)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2516
  by (import hollight DEF_MEM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2517
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2518
definition ALL2 :: "('q_16997 => 'q_17004 => bool)
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2519
=> 'q_16997 hollight.list => 'q_17004 hollight.list => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2520
  "ALL2 ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2521
SOME ALL2::('q_16997::type => 'q_17004::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2522
           => 'q_16997::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2523
              => 'q_17004::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2524
   (ALL (P::'q_16997::type => 'q_17004::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2525
       l2::'q_17004::type hollight.list. ALL2 P NIL l2 = (l2 = NIL)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2526
   (ALL (h1::'q_16997::type) (P::'q_16997::type => 'q_17004::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2527
       (t1::'q_16997::type hollight.list) l2::'q_17004::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2528
       ALL2 P (CONS h1 t1) l2 =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2529
       COND (l2 = NIL) False (P h1 (HD l2) & ALL2 P t1 (TL l2)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2530
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2531
lemma DEF_ALL2: "ALL2 =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2532
(SOME ALL2::('q_16997::type => 'q_17004::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2533
            => 'q_16997::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2534
               => 'q_17004::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2535
    (ALL (P::'q_16997::type => 'q_17004::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2536
        l2::'q_17004::type hollight.list. ALL2 P NIL l2 = (l2 = NIL)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2537
    (ALL (h1::'q_16997::type) (P::'q_16997::type => 'q_17004::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2538
        (t1::'q_16997::type hollight.list) l2::'q_17004::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2539
        ALL2 P (CONS h1 t1) l2 =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2540
        COND (l2 = NIL) False (P h1 (HD l2) & ALL2 P t1 (TL l2))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2541
  by (import hollight DEF_ALL2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2542
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2543
lemma ALL2: "ALL2 (P::'q_17059::type => 'q_17058::type => bool) NIL NIL = True &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2544
ALL2 P (CONS (h1::'q_17059::type) (t1::'q_17059::type hollight.list)) NIL =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2545
False &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2546
ALL2 P NIL (CONS (h2::'q_17058::type) (t2::'q_17058::type hollight.list)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2547
False &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2548
ALL2 P (CONS h1 t1) (CONS h2 t2) = (P h1 h2 & ALL2 P t1 t2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2549
  by (import hollight ALL2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2550
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2551
definition MAP2 :: "('q_17089 => 'q_17096 => 'q_17086)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2552
=> 'q_17089 hollight.list
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2553
   => 'q_17096 hollight.list => 'q_17086 hollight.list" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2554
  "MAP2 ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2555
SOME MAP2::('q_17089::type => 'q_17096::type => 'q_17086::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2556
           => 'q_17089::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2557
              => 'q_17096::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2558
                 => 'q_17086::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2559
   (ALL (f::'q_17089::type => 'q_17096::type => 'q_17086::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2560
       l::'q_17096::type hollight.list. MAP2 f NIL l = NIL) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2561
   (ALL (h1::'q_17089::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2562
       (f::'q_17089::type => 'q_17096::type => 'q_17086::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2563
       (t1::'q_17089::type hollight.list) l::'q_17096::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2564
       MAP2 f (CONS h1 t1) l = CONS (f h1 (HD l)) (MAP2 f t1 (TL l)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2565
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2566
lemma DEF_MAP2: "MAP2 =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2567
(SOME MAP2::('q_17089::type => 'q_17096::type => 'q_17086::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2568
            => 'q_17089::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2569
               => 'q_17096::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2570
                  => 'q_17086::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2571
    (ALL (f::'q_17089::type => 'q_17096::type => 'q_17086::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2572
        l::'q_17096::type hollight.list. MAP2 f NIL l = NIL) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2573
    (ALL (h1::'q_17089::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2574
        (f::'q_17089::type => 'q_17096::type => 'q_17086::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2575
        (t1::'q_17089::type hollight.list) l::'q_17096::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2576
        MAP2 f (CONS h1 t1) l = CONS (f h1 (HD l)) (MAP2 f t1 (TL l))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2577
  by (import hollight DEF_MAP2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2578
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2579
lemma MAP2: "MAP2 (f::'q_17131::type => 'q_17130::type => 'q_17137::type) NIL NIL = NIL &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2580
MAP2 f (CONS (h1::'q_17131::type) (t1::'q_17131::type hollight.list))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2581
 (CONS (h2::'q_17130::type) (t2::'q_17130::type hollight.list)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2582
CONS (f h1 h2) (MAP2 f t1 t2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2583
  by (import hollight MAP2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2584
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2585
definition EL :: "nat => 'q_17157 hollight.list => 'q_17157" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2586
  "EL ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2587
SOME EL::nat => 'q_17157::type hollight.list => 'q_17157::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2588
   (ALL l::'q_17157::type hollight.list. EL 0 l = HD l) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2589
   (ALL (n::nat) l::'q_17157::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2590
       EL (Suc n) l = EL n (TL l))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2591
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2592
lemma DEF_EL: "EL =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2593
(SOME EL::nat => 'q_17157::type hollight.list => 'q_17157::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2594
    (ALL l::'q_17157::type hollight.list. EL 0 l = HD l) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2595
    (ALL (n::nat) l::'q_17157::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2596
        EL (Suc n) l = EL n (TL l)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2597
  by (import hollight DEF_EL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2598
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2599
definition FILTER :: "('q_17182 => bool) => 'q_17182 hollight.list => 'q_17182 hollight.list" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2600
  "FILTER ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2601
SOME FILTER::('q_17182::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2602
             => 'q_17182::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2603
                => 'q_17182::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2604
   (ALL P::'q_17182::type => bool. FILTER P NIL = NIL) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2605
   (ALL (h::'q_17182::type) (P::'q_17182::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2606
       t::'q_17182::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2607
       FILTER P (CONS h t) = COND (P h) (CONS h (FILTER P t)) (FILTER P t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2608
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2609
lemma DEF_FILTER: "FILTER =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2610
(SOME FILTER::('q_17182::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2611
              => 'q_17182::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2612
                 => 'q_17182::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2613
    (ALL P::'q_17182::type => bool. FILTER P NIL = NIL) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2614
    (ALL (h::'q_17182::type) (P::'q_17182::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2615
        t::'q_17182::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2616
        FILTER P (CONS h t) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2617
        COND (P h) (CONS h (FILTER P t)) (FILTER P t)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2618
  by (import hollight DEF_FILTER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2619
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2620
definition ASSOC :: "'q_17211 => ('q_17211 * 'q_17205) hollight.list => 'q_17205" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2621
  "ASSOC ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2622
SOME ASSOC::'q_17211::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2623
            => ('q_17211::type * 'q_17205::type) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2624
               => 'q_17205::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2625
   ALL (h::'q_17211::type * 'q_17205::type) (a::'q_17211::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2626
      t::('q_17211::type * 'q_17205::type) hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2627
      ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2628
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2629
lemma DEF_ASSOC: "ASSOC =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2630
(SOME ASSOC::'q_17211::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2631
             => ('q_17211::type * 'q_17205::type) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2632
                => 'q_17205::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2633
    ALL (h::'q_17211::type * 'q_17205::type) (a::'q_17211::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2634
       t::('q_17211::type * 'q_17205::type) hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2635
       ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2636
  by (import hollight DEF_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2637
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2638
definition ITLIST2 :: "('q_17235 => 'q_17243 => 'q_17233 => 'q_17233)
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2639
=> 'q_17235 hollight.list => 'q_17243 hollight.list => 'q_17233 => 'q_17233" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2640
  "ITLIST2 ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2641
SOME ITLIST2::('q_17235::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2642
               => 'q_17243::type => 'q_17233::type => 'q_17233::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2643
              => 'q_17235::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2644
                 => 'q_17243::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2645
                    => 'q_17233::type => 'q_17233::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2646
   (ALL (f::'q_17235::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2647
            => 'q_17243::type => 'q_17233::type => 'q_17233::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2648
       (l2::'q_17243::type hollight.list) b::'q_17233::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2649
       ITLIST2 f NIL l2 b = b) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2650
   (ALL (h1::'q_17235::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2651
       (f::'q_17235::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2652
           => 'q_17243::type => 'q_17233::type => 'q_17233::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2653
       (t1::'q_17235::type hollight.list) (l2::'q_17243::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2654
       b::'q_17233::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2655
       ITLIST2 f (CONS h1 t1) l2 b = f h1 (HD l2) (ITLIST2 f t1 (TL l2) b))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2656
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2657
lemma DEF_ITLIST2: "ITLIST2 =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2658
(SOME ITLIST2::('q_17235::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2659
                => 'q_17243::type => 'q_17233::type => 'q_17233::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2660
               => 'q_17235::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2661
                  => 'q_17243::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2662
                     => 'q_17233::type => 'q_17233::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2663
    (ALL (f::'q_17235::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2664
             => 'q_17243::type => 'q_17233::type => 'q_17233::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2665
        (l2::'q_17243::type hollight.list) b::'q_17233::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2666
        ITLIST2 f NIL l2 b = b) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2667
    (ALL (h1::'q_17235::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2668
        (f::'q_17235::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2669
            => 'q_17243::type => 'q_17233::type => 'q_17233::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2670
        (t1::'q_17235::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2671
        (l2::'q_17243::type hollight.list) b::'q_17233::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2672
        ITLIST2 f (CONS h1 t1) l2 b =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2673
        f h1 (HD l2) (ITLIST2 f t1 (TL l2) b)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2674
  by (import hollight DEF_ITLIST2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2675
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2676
lemma ITLIST2: "ITLIST2
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2677
 (f::'q_17277::type => 'q_17276::type => 'q_17275::type => 'q_17275::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2678
 NIL NIL (b::'q_17275::type) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2679
b &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2680
ITLIST2 f (CONS (h1::'q_17277::type) (t1::'q_17277::type hollight.list))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2681
 (CONS (h2::'q_17276::type) (t2::'q_17276::type hollight.list)) b =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2682
f h1 h2 (ITLIST2 f t1 t2 b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2683
  by (import hollight ITLIST2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2684
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2685
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2686
  ZIP :: "'q_17307 hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2687
=> 'q_17315 hollight.list => ('q_17307 * 'q_17315) hollight.list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2688
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2689
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2690
  ZIP_def: "hollight.ZIP ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2691
SOME ZIP::'q_17307::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2692
          => 'q_17315::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2693
             => ('q_17307::type * 'q_17315::type) hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2694
   (ALL l2::'q_17315::type hollight.list. ZIP NIL l2 = NIL) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2695
   (ALL (h1::'q_17307::type) (t1::'q_17307::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2696
       l2::'q_17315::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2697
       ZIP (CONS h1 t1) l2 = CONS (h1, HD l2) (ZIP t1 (TL l2)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2698
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2699
lemma DEF_ZIP: "hollight.ZIP =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2700
(SOME ZIP::'q_17307::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2701
           => 'q_17315::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2702
              => ('q_17307::type * 'q_17315::type) hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2703
    (ALL l2::'q_17315::type hollight.list. ZIP NIL l2 = NIL) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2704
    (ALL (h1::'q_17307::type) (t1::'q_17307::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2705
        l2::'q_17315::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2706
        ZIP (CONS h1 t1) l2 = CONS (h1, HD l2) (ZIP t1 (TL l2))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2707
  by (import hollight DEF_ZIP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2708
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2709
lemma ZIP: "(op &::bool => bool => bool)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2710
 ((op =::('q_17326::type * 'q_17327::type) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2711
         => ('q_17326::type * 'q_17327::type) hollight.list => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2712
   ((hollight.ZIP::'q_17326::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2713
                   => 'q_17327::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2714
                      => ('q_17326::type * 'q_17327::type) hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2715
     (NIL::'q_17326::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2716
     (NIL::'q_17327::type hollight.list))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2717
   (NIL::('q_17326::type * 'q_17327::type) hollight.list))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2718
 ((op =::('q_17351::type * 'q_17352::type) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2719
         => ('q_17351::type * 'q_17352::type) hollight.list => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2720
   ((hollight.ZIP::'q_17351::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2721
                   => 'q_17352::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2722
                      => ('q_17351::type * 'q_17352::type) hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2723
     ((CONS::'q_17351::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2724
             => 'q_17351::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2725
                => 'q_17351::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2726
       (h1::'q_17351::type) (t1::'q_17351::type hollight.list))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2727
     ((CONS::'q_17352::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2728
             => 'q_17352::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2729
                => 'q_17352::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2730
       (h2::'q_17352::type) (t2::'q_17352::type hollight.list)))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2731
   ((CONS::'q_17351::type * 'q_17352::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2732
           => ('q_17351::type * 'q_17352::type) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2733
              => ('q_17351::type * 'q_17352::type) hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2734
     ((Pair::'q_17351::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2735
             => 'q_17352::type => 'q_17351::type * 'q_17352::type)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2736
       h1 h2)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2737
     ((hollight.ZIP::'q_17351::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2738
                     => 'q_17352::type hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2739
                        => ('q_17351::type * 'q_17352::type) hollight.list)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2740
       t1 t2)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2741
  by (import hollight ZIP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2742
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2743
lemma NOT_CONS_NIL: "ALL (x::'A::type) xa::'A::type hollight.list. CONS x xa ~= NIL"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2744
  by (import hollight NOT_CONS_NIL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2745
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2746
lemma LAST_CLAUSES: "LAST (CONS (h::'A::type) NIL) = h &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2747
LAST (CONS h (CONS (k::'A::type) (t::'A::type hollight.list))) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2748
LAST (CONS k t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2749
  by (import hollight LAST_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2750
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2751
lemma APPEND_NIL: "ALL l::'A::type hollight.list. APPEND l NIL = l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2752
  by (import hollight APPEND_NIL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2753
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2754
lemma APPEND_ASSOC: "ALL (l::'A::type hollight.list) (m::'A::type hollight.list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2755
   n::'A::type hollight.list. APPEND l (APPEND m n) = APPEND (APPEND l m) n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2756
  by (import hollight APPEND_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2757
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2758
lemma REVERSE_APPEND: "ALL (l::'A::type hollight.list) m::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2759
   REVERSE (APPEND l m) = APPEND (REVERSE m) (REVERSE l)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2760
  by (import hollight REVERSE_APPEND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2761
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2762
lemma REVERSE_REVERSE: "ALL l::'A::type hollight.list. REVERSE (REVERSE l) = l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2763
  by (import hollight REVERSE_REVERSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2764
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2765
lemma CONS_11: "ALL (x::'A::type) (xa::'A::type) (xb::'A::type hollight.list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2766
   xc::'A::type hollight.list. (CONS x xb = CONS xa xc) = (x = xa & xb = xc)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2767
  by (import hollight CONS_11)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2768
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2769
lemma list_CASES: "ALL l::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2770
   l = NIL | (EX (h::'A::type) t::'A::type hollight.list. l = CONS h t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2771
  by (import hollight list_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2772
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2773
lemma LENGTH_APPEND: "ALL (l::'A::type hollight.list) m::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2774
   LENGTH (APPEND l m) = LENGTH l + LENGTH m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2775
  by (import hollight LENGTH_APPEND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2776
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2777
lemma MAP_APPEND: "ALL (f::'A::type => 'B::type) (l1::'A::type hollight.list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2778
   l2::'A::type hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2779
   MAP f (APPEND l1 l2) = APPEND (MAP f l1) (MAP f l2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2780
  by (import hollight MAP_APPEND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2781
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2782
lemma LENGTH_MAP: "ALL (l::'A::type hollight.list) f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2783
   LENGTH (MAP f l) = LENGTH l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2784
  by (import hollight LENGTH_MAP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2785
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2786
lemma LENGTH_EQ_NIL: "ALL l::'A::type hollight.list. (LENGTH l = 0) = (l = NIL)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2787
  by (import hollight LENGTH_EQ_NIL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2788
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2789
lemma LENGTH_EQ_CONS: "ALL (l::'q_17659::type hollight.list) n::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2790
   (LENGTH l = Suc n) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2791
   (EX (h::'q_17659::type) t::'q_17659::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2792
       l = CONS h t & LENGTH t = n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2793
  by (import hollight LENGTH_EQ_CONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2794
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2795
lemma MAP_o: "ALL (f::'A::type => 'B::type) (g::'B::type => 'C::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2796
   l::'A::type hollight.list. MAP (g o f) l = MAP g (MAP f l)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2797
  by (import hollight MAP_o)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2798
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2799
lemma MAP_EQ: "ALL (f::'q_17723::type => 'q_17734::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2800
   (g::'q_17723::type => 'q_17734::type) l::'q_17723::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2801
   ALL_list (%x::'q_17723::type. f x = g x) l --> MAP f l = MAP g l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2802
  by (import hollight MAP_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2803
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2804
lemma ALL_IMP: "ALL (P::'q_17764::type => bool) (Q::'q_17764::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2805
   l::'q_17764::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2806
   (ALL x::'q_17764::type. MEM x l & P x --> Q x) & ALL_list P l -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2807
   ALL_list Q l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2808
  by (import hollight ALL_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2809
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2810
lemma NOT_EX: "ALL (P::'q_17792::type => bool) l::'q_17792::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2811
   (~ EX P l) = ALL_list (%x::'q_17792::type. ~ P x) l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2812
  by (import hollight NOT_EX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2813
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2814
lemma NOT_ALL: "ALL (P::'q_17814::type => bool) l::'q_17814::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2815
   (~ ALL_list P l) = EX (%x::'q_17814::type. ~ P x) l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2816
  by (import hollight NOT_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2817
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2818
lemma ALL_MAP: "ALL (P::'q_17836::type => bool) (f::'q_17835::type => 'q_17836::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2819
   l::'q_17835::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2820
   ALL_list P (MAP f l) = ALL_list (P o f) l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2821
  by (import hollight ALL_MAP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2822
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2823
lemma ALL_T: "All (ALL_list (%x::'q_17854::type. True))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2824
  by (import hollight ALL_T)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2825
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2826
lemma MAP_EQ_ALL2: "ALL (l::'q_17879::type hollight.list) m::'q_17879::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2827
   ALL2
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2828
    (%(x::'q_17879::type) y::'q_17879::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2829
        (f::'q_17879::type => 'q_17890::type) x = f y)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2830
    l m -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2831
   MAP f l = MAP f m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2832
  by (import hollight MAP_EQ_ALL2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2833
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2834
lemma ALL2_MAP: "ALL (P::'q_17921::type => 'q_17922::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2835
   (f::'q_17922::type => 'q_17921::type) l::'q_17922::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2836
   ALL2 P (MAP f l) l = ALL_list (%a::'q_17922::type. P (f a) a) l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2837
  by (import hollight ALL2_MAP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2838
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2839
lemma MAP_EQ_DEGEN: "ALL (l::'q_17939::type hollight.list) f::'q_17939::type => 'q_17939::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2840
   ALL_list (%x::'q_17939::type. f x = x) l --> MAP f l = l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2841
  by (import hollight MAP_EQ_DEGEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2842
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2843
lemma ALL2_AND_RIGHT: "ALL (l::'q_17982::type hollight.list) (m::'q_17981::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2844
   (P::'q_17982::type => bool) Q::'q_17982::type => 'q_17981::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2845
   ALL2 (%(x::'q_17982::type) y::'q_17981::type. P x & Q x y) l m =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2846
   (ALL_list P l & ALL2 Q l m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2847
  by (import hollight ALL2_AND_RIGHT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2848
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2849
lemma ITLIST_EXTRA: "ALL l::'q_18019::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2850
   ITLIST (f::'q_18019::type => 'q_18018::type => 'q_18018::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2851
    (APPEND l (CONS (a::'q_18019::type) NIL)) (b::'q_18018::type) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2852
   ITLIST f l (f a b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2853
  by (import hollight ITLIST_EXTRA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2854
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2855
lemma ALL_MP: "ALL (P::'q_18045::type => bool) (Q::'q_18045::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2856
   l::'q_18045::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2857
   ALL_list (%x::'q_18045::type. P x --> Q x) l & ALL_list P l -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2858
   ALL_list Q l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2859
  by (import hollight ALL_MP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2860
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2861
lemma AND_ALL: "ALL x::'q_18075::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2862
   (ALL_list (P::'q_18075::type => bool) x &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2863
    ALL_list (Q::'q_18075::type => bool) x) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2864
   ALL_list (%x::'q_18075::type. P x & Q x) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2865
  by (import hollight AND_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2866
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2867
lemma EX_IMP: "ALL (P::'q_18105::type => bool) (Q::'q_18105::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2868
   l::'q_18105::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2869
   (ALL x::'q_18105::type. MEM x l & P x --> Q x) & EX P l --> EX Q l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2870
  by (import hollight EX_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2871
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2872
lemma ALL_MEM: "ALL (P::'q_18132::type => bool) l::'q_18132::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2873
   (ALL x::'q_18132::type. MEM x l --> P x) = ALL_list P l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2874
  by (import hollight ALL_MEM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2875
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2876
lemma LENGTH_REPLICATE: "ALL (n::nat) x::'q_18150::type. LENGTH (REPLICATE n x) = n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2877
  by (import hollight LENGTH_REPLICATE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2878
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2879
lemma EX_MAP: "ALL (P::'q_18174::type => bool) (f::'q_18173::type => 'q_18174::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2880
   l::'q_18173::type hollight.list. EX P (MAP f l) = EX (P o f) l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2881
  by (import hollight EX_MAP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2882
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2883
lemma EXISTS_EX: "ALL (P::'q_18212::type => 'q_18211::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2884
   l::'q_18211::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2885
   (EX x::'q_18212::type. EX (P x) l) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2886
   EX (%s::'q_18211::type. EX x::'q_18212::type. P x s) l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2887
  by (import hollight EXISTS_EX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2888
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2889
lemma FORALL_ALL: "ALL (P::'q_18242::type => 'q_18241::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2890
   l::'q_18241::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2891
   (ALL x::'q_18242::type. ALL_list (P x) l) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2892
   ALL_list (%s::'q_18241::type. ALL x::'q_18242::type. P x s) l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2893
  by (import hollight FORALL_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2894
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2895
lemma MEM_APPEND: "ALL (x::'q_18270::type) (l1::'q_18270::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2896
   l2::'q_18270::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2897
   MEM x (APPEND l1 l2) = (MEM x l1 | MEM x l2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2898
  by (import hollight MEM_APPEND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2899
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2900
lemma MEM_MAP: "ALL (f::'q_18306::type => 'q_18303::type) (y::'q_18303::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2901
   l::'q_18306::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2902
   MEM y (MAP f l) = (EX x::'q_18306::type. MEM x l & y = f x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2903
  by (import hollight MEM_MAP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2904
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2905
lemma FILTER_APPEND: "ALL (P::'q_18337::type => bool) (l1::'q_18337::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2906
   l2::'q_18337::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2907
   FILTER P (APPEND l1 l2) = APPEND (FILTER P l1) (FILTER P l2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2908
  by (import hollight FILTER_APPEND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2909
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2910
lemma FILTER_MAP: "ALL (P::'q_18364::type => bool) (f::'q_18371::type => 'q_18364::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2911
   l::'q_18371::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2912
   FILTER P (MAP f l) = MAP f (FILTER (P o f) l)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2913
  by (import hollight FILTER_MAP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2914
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2915
lemma MEM_FILTER: "ALL (P::'q_18398::type => bool) (l::'q_18398::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2916
   x::'q_18398::type. MEM x (FILTER P l) = (P x & MEM x l)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2917
  by (import hollight MEM_FILTER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2918
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2919
lemma EX_MEM: "ALL (P::'q_18419::type => bool) l::'q_18419::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2920
   (EX x::'q_18419::type. P x & MEM x l) = EX P l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2921
  by (import hollight EX_MEM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2922
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2923
lemma MAP_FST_ZIP: "ALL (l1::'q_18439::type hollight.list) l2::'q_18441::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2924
   LENGTH l1 = LENGTH l2 --> MAP fst (hollight.ZIP l1 l2) = l1"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2925
  by (import hollight MAP_FST_ZIP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2926
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2927
lemma MAP_SND_ZIP: "ALL (l1::'q_18470::type hollight.list) l2::'q_18472::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2928
   LENGTH l1 = LENGTH l2 --> MAP snd (hollight.ZIP l1 l2) = l2"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2929
  by (import hollight MAP_SND_ZIP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2930
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2931
lemma MEM_ASSOC: "ALL (l::('q_18516::type * 'q_18500::type) hollight.list) x::'q_18516::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2932
   MEM (x, ASSOC x l) l = MEM x (MAP fst l)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2933
  by (import hollight MEM_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2934
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2935
lemma ALL_APPEND: "ALL (P::'q_18537::type => bool) (l1::'q_18537::type hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2936
   l2::'q_18537::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2937
   ALL_list P (APPEND l1 l2) = (ALL_list P l1 & ALL_list P l2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2938
  by (import hollight ALL_APPEND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2939
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2940
lemma MEM_EL: "ALL (l::'q_18560::type hollight.list) n::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2941
   < n (LENGTH l) --> MEM (EL n l) l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2942
  by (import hollight MEM_EL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2943
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2944
lemma ALL2_MAP2: "ALL (l::'q_18603::type hollight.list) m::'q_18604::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2945
   ALL2 (P::'q_18602::type => 'q_18601::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2946
    (MAP (f::'q_18603::type => 'q_18602::type) l)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2947
    (MAP (g::'q_18604::type => 'q_18601::type) m) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2948
   ALL2 (%(x::'q_18603::type) y::'q_18604::type. P (f x) (g y)) l m"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2949
  by (import hollight ALL2_MAP2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2950
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2951
lemma AND_ALL2: "ALL (P::'q_18650::type => 'q_18649::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2952
   (Q::'q_18650::type => 'q_18649::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2953
   (x::'q_18650::type hollight.list) xa::'q_18649::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2954
   (ALL2 P x xa & ALL2 Q x xa) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2955
   ALL2 (%(x::'q_18650::type) y::'q_18649::type. P x y & Q x y) x xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2956
  by (import hollight AND_ALL2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2957
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2958
lemma ALL2_ALL: "ALL (P::'q_18672::type => 'q_18672::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2959
   l::'q_18672::type hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2960
   ALL2 P l l = ALL_list (%x::'q_18672::type. P x x) l"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2961
  by (import hollight ALL2_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2962
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2963
lemma APPEND_EQ_NIL: "ALL (x::'q_18701::type hollight.list) xa::'q_18701::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2964
   (APPEND x xa = NIL) = (x = NIL & xa = NIL)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2965
  by (import hollight APPEND_EQ_NIL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2966
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2967
lemma LENGTH_MAP2: "ALL (f::'q_18721::type => 'q_18723::type => 'q_18734::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  2968
   (l::'q_18721::type hollight.list) m::'q_18723::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2969
   LENGTH l = LENGTH m --> LENGTH (MAP2 f l m) = LENGTH m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2970
  by (import hollight LENGTH_MAP2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2971
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2972
lemma MONO_ALL: "(ALL x::'A::type. (P::'A::type => bool) x --> (Q::'A::type => bool) x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2973
ALL_list P (l::'A::type hollight.list) --> ALL_list Q l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2974
  by (import hollight MONO_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2975
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2976
lemma MONO_ALL2: "(ALL (x::'A::type) y::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2977
    (P::'A::type => 'B::type => bool) x y -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2978
    (Q::'A::type => 'B::type => bool) x y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2979
ALL2 P (l::'A::type hollight.list) (l'::'B::type hollight.list) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2980
ALL2 Q l l'"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2981
  by (import hollight MONO_ALL2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2982
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  2983
definition dist :: "nat * nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2984
  "dist == %u::nat * nat. fst u - snd u + (snd u - fst u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2985
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2986
lemma DEF_dist: "dist = (%u::nat * nat. fst u - snd u + (snd u - fst u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2987
  by (import hollight DEF_dist)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2988
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2989
lemma DIST_REFL: "ALL x::nat. dist (x, x) = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2990
  by (import hollight DIST_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2991
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2992
lemma DIST_LZERO: "ALL x::nat. dist (0, x) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2993
  by (import hollight DIST_LZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2994
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2995
lemma DIST_RZERO: "ALL x::nat. dist (x, 0) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2996
  by (import hollight DIST_RZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2997
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2998
lemma DIST_SYM: "ALL (x::nat) xa::nat. dist (x, xa) = dist (xa, x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  2999
  by (import hollight DIST_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3000
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3001
lemma DIST_LADD: "ALL (x::nat) (xa::nat) xb::nat. dist (x + xb, x + xa) = dist (xb, xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3002
  by (import hollight DIST_LADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3003
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3004
lemma DIST_RADD: "ALL (x::nat) (xa::nat) xb::nat. dist (x + xa, xb + xa) = dist (x, xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3005
  by (import hollight DIST_RADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3006
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3007
lemma DIST_LADD_0: "ALL (x::nat) xa::nat. dist (x + xa, x) = xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3008
  by (import hollight DIST_LADD_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3009
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3010
lemma DIST_RADD_0: "ALL (x::nat) xa::nat. dist (x, x + xa) = xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3011
  by (import hollight DIST_RADD_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3012
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3013
lemma DIST_LMUL: "ALL (x::nat) (xa::nat) xb::nat. x * dist (xa, xb) = dist (x * xa, x * xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3014
  by (import hollight DIST_LMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3015
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3016
lemma DIST_RMUL: "ALL (x::nat) (xa::nat) xb::nat. dist (x, xa) * xb = dist (x * xb, xa * xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3017
  by (import hollight DIST_RMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3018
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3019
lemma DIST_EQ_0: "ALL (x::nat) xa::nat. (dist (x, xa) = 0) = (x = xa)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3020
  by (import hollight DIST_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3021
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3022
lemma DIST_ELIM_THM: "(P::nat => bool) (dist (x::nat, y::nat)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3023
(ALL d::nat. (x = y + d --> P d) & (y = x + d --> P d))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3024
  by (import hollight DIST_ELIM_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3025
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3026
lemma DIST_LE_CASES: "ALL (m::nat) (n::nat) p::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3027
   <= (dist (m, n)) p = (<= m (n + p) & <= n (m + p))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3028
  by (import hollight DIST_LE_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3029
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3030
lemma DIST_ADDBOUND: "ALL (m::nat) n::nat. <= (dist (m, n)) (m + n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3031
  by (import hollight DIST_ADDBOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3032
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3033
lemma DIST_TRIANGLE: "ALL (m::nat) (n::nat) p::nat. <= (dist (m, p)) (dist (m, n) + dist (n, p))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3034
  by (import hollight DIST_TRIANGLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3035
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3036
lemma DIST_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3037
   <= (dist (m + n, p + q)) (dist (m, p) + dist (n, q))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3038
  by (import hollight DIST_ADD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3039
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3040
lemma DIST_ADD2_REV: "ALL (m::nat) (n::nat) (p::nat) q::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3041
   <= (dist (m, p)) (dist (m + n, p + q) + dist (n, q))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3042
  by (import hollight DIST_ADD2_REV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3043
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3044
lemma DIST_TRIANGLE_LE: "ALL (m::nat) (n::nat) (p::nat) q::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3045
   <= (dist (m, n) + dist (n, p)) q --> <= (dist (m, p)) q"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3046
  by (import hollight DIST_TRIANGLE_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3047
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3048
lemma DIST_TRIANGLES_LE: "ALL (m::nat) (n::nat) (p::nat) (q::nat) (r::nat) s::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3049
   <= (dist (m, n)) r & <= (dist (p, q)) s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3050
   <= (dist (m, p)) (dist (n, q) + (r + s))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3051
  by (import hollight DIST_TRIANGLES_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3052
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3053
lemma BOUNDS_LINEAR: "ALL (A::nat) (B::nat) C::nat. (ALL n::nat. <= (A * n) (B * n + C)) = <= A B"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3054
  by (import hollight BOUNDS_LINEAR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3055
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3056
lemma BOUNDS_LINEAR_0: "ALL (A::nat) B::nat. (ALL n::nat. <= (A * n) B) = (A = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3057
  by (import hollight BOUNDS_LINEAR_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3058
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3059
lemma BOUNDS_DIVIDED: "ALL P::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3060
   (EX B::nat. ALL n::nat. <= (P n) B) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3061
   (EX (x::nat) B::nat. ALL n::nat. <= (n * P n) (x * n + B))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3062
  by (import hollight BOUNDS_DIVIDED)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3063
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3064
lemma BOUNDS_NOTZERO: "ALL (P::nat => nat => nat) (A::nat) B::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3065
   P 0 0 = 0 & (ALL (m::nat) n::nat. <= (P m n) (A * (m + n) + B)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3066
   (EX x::nat. ALL (m::nat) n::nat. <= (P m n) (x * (m + n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3067
  by (import hollight BOUNDS_NOTZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3068
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3069
lemma BOUNDS_IGNORE: "ALL (P::nat => nat) Q::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3070
   (EX B::nat. ALL i::nat. <= (P i) (Q i + B)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3071
   (EX (x::nat) N::nat. ALL i::nat. <= N i --> <= (P i) (Q i + x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3072
  by (import hollight BOUNDS_IGNORE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3073
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3074
definition is_nadd :: "(nat => nat) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3075
  "is_nadd ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3076
%u::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3077
   EX B::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3078
      ALL (m::nat) n::nat. <= (dist (m * u n, n * u m)) (B * (m + n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3079
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3080
lemma DEF_is_nadd: "is_nadd =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3081
(%u::nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3082
    EX B::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3083
       ALL (m::nat) n::nat. <= (dist (m * u n, n * u m)) (B * (m + n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3084
  by (import hollight DEF_is_nadd)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3085
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3086
lemma is_nadd_0: "is_nadd (%n::nat. 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3087
  by (import hollight is_nadd_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3088
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3089
typedef (open) nadd = "Collect is_nadd"  morphisms "dest_nadd" "mk_nadd"
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3090
  apply (rule light_ex_imp_nonempty[where t="%n::nat. NUMERAL 0"])
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3091
  by (import hollight TYDEF_nadd)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3092
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3093
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3094
  dest_nadd :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3095
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3096
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3097
  mk_nadd :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3098
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3099
lemmas "TYDEF_nadd_@intern" = typedef_hol2hollight 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3100
  [where a="a :: nadd" and r=r ,
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3101
   OF type_definition_nadd]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3102
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3103
lemma NADD_CAUCHY: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3104
   EX xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3105
      ALL (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3106
         <= (dist (xb * dest_nadd x xc, xc * dest_nadd x xb))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3107
          (xa * (xb + xc))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3108
  by (import hollight NADD_CAUCHY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3109
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3110
lemma NADD_BOUND: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3111
   EX (xa::nat) B::nat. ALL n::nat. <= (dest_nadd x n) (xa * n + B)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3112
  by (import hollight NADD_BOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3113
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3114
lemma NADD_MULTIPLICATIVE: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3115
   EX xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3116
      ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3117
         <= (dist (dest_nadd x (m * n), m * dest_nadd x n)) (xa * m + xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3118
  by (import hollight NADD_MULTIPLICATIVE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3119
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3120
lemma NADD_ADDITIVE: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3121
   EX xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3122
      ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3123
         <= (dist (dest_nadd x (m + n), dest_nadd x m + dest_nadd x n)) xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3124
  by (import hollight NADD_ADDITIVE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3125
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3126
lemma NADD_SUC: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3127
   EX xa::nat. ALL n::nat. <= (dist (dest_nadd x (Suc n), dest_nadd x n)) xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3128
  by (import hollight NADD_SUC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3129
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3130
lemma NADD_DIST_LEMMA: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3131
   EX xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3132
      ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3133
         <= (dist (dest_nadd x (m + n), dest_nadd x m)) (xa * n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3134
  by (import hollight NADD_DIST_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3135
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3136
lemma NADD_DIST: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3137
   EX xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3138
      ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3139
         <= (dist (dest_nadd x m, dest_nadd x n)) (xa * dist (m, n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3140
  by (import hollight NADD_DIST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3141
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3142
lemma NADD_ALTMUL: "ALL (x::nadd) y::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3143
   EX (A::nat) B::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3144
      ALL n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3145
         <= (dist
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3146
              (n * dest_nadd x (dest_nadd y n),
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3147
               dest_nadd x n * dest_nadd y n))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3148
          (A * n + B)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3149
  by (import hollight NADD_ALTMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3150
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3151
definition nadd_eq :: "nadd => nadd => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3152
  "nadd_eq ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3153
%(u::nadd) ua::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3154
   EX B::nat. ALL n::nat. <= (dist (dest_nadd u n, dest_nadd ua n)) B"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3155
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3156
lemma DEF_nadd_eq: "nadd_eq =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3157
(%(u::nadd) ua::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3158
    EX B::nat. ALL n::nat. <= (dist (dest_nadd u n, dest_nadd ua n)) B)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3159
  by (import hollight DEF_nadd_eq)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3160
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3161
lemma NADD_EQ_REFL: "ALL x::nadd. nadd_eq x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3162
  by (import hollight NADD_EQ_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3163
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3164
lemma NADD_EQ_SYM: "ALL (x::nadd) y::nadd. nadd_eq x y = nadd_eq y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3165
  by (import hollight NADD_EQ_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3166
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3167
lemma NADD_EQ_TRANS: "ALL (x::nadd) (y::nadd) z::nadd. nadd_eq x y & nadd_eq y z --> nadd_eq x z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3168
  by (import hollight NADD_EQ_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3169
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3170
definition nadd_of_num :: "nat => nadd" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3171
  "nadd_of_num == %u::nat. mk_nadd (op * u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3172
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3173
lemma DEF_nadd_of_num: "nadd_of_num = (%u::nat. mk_nadd (op * u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3174
  by (import hollight DEF_nadd_of_num)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3175
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3176
lemma NADD_OF_NUM: "ALL x::nat. dest_nadd (nadd_of_num x) = op * x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3177
  by (import hollight NADD_OF_NUM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3178
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3179
lemma NADD_OF_NUM_WELLDEF: "ALL (m::nat) n::nat. m = n --> nadd_eq (nadd_of_num m) (nadd_of_num n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3180
  by (import hollight NADD_OF_NUM_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3181
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3182
lemma NADD_OF_NUM_EQ: "ALL (m::nat) n::nat. nadd_eq (nadd_of_num m) (nadd_of_num n) = (m = n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3183
  by (import hollight NADD_OF_NUM_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3184
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3185
definition nadd_le :: "nadd => nadd => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3186
  "nadd_le ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3187
%(u::nadd) ua::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3188
   EX B::nat. ALL n::nat. <= (dest_nadd u n) (dest_nadd ua n + B)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3189
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3190
lemma DEF_nadd_le: "nadd_le =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3191
(%(u::nadd) ua::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3192
    EX B::nat. ALL n::nat. <= (dest_nadd u n) (dest_nadd ua n + B))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3193
  by (import hollight DEF_nadd_le)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3194
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3195
lemma NADD_LE_WELLDEF_LEMMA: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3196
   nadd_eq x x' & nadd_eq y y' & nadd_le x y --> nadd_le x' y'"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3197
  by (import hollight NADD_LE_WELLDEF_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3198
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3199
lemma NADD_LE_WELLDEF: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3200
   nadd_eq x x' & nadd_eq y y' --> nadd_le x y = nadd_le x' y'"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3201
  by (import hollight NADD_LE_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3202
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3203
lemma NADD_LE_REFL: "ALL x::nadd. nadd_le x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3204
  by (import hollight NADD_LE_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3205
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3206
lemma NADD_LE_TRANS: "ALL (x::nadd) (y::nadd) z::nadd. nadd_le x y & nadd_le y z --> nadd_le x z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3207
  by (import hollight NADD_LE_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3208
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3209
lemma NADD_LE_ANTISYM: "ALL (x::nadd) y::nadd. (nadd_le x y & nadd_le y x) = nadd_eq x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3210
  by (import hollight NADD_LE_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3211
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3212
lemma NADD_LE_TOTAL_LEMMA: "ALL (x::nadd) y::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3213
   ~ nadd_le x y -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3214
   (ALL B::nat. EX n::nat. n ~= 0 & < (dest_nadd y n + B) (dest_nadd x n))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3215
  by (import hollight NADD_LE_TOTAL_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3216
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3217
lemma NADD_LE_TOTAL: "ALL (x::nadd) y::nadd. nadd_le x y | nadd_le y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3218
  by (import hollight NADD_LE_TOTAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3219
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3220
lemma NADD_ARCH: "ALL x::nadd. EX xa::nat. nadd_le x (nadd_of_num xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3221
  by (import hollight NADD_ARCH)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3222
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3223
lemma NADD_OF_NUM_LE: "ALL (m::nat) n::nat. nadd_le (nadd_of_num m) (nadd_of_num n) = <= m n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3224
  by (import hollight NADD_OF_NUM_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3225
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3226
definition nadd_add :: "nadd => nadd => nadd" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3227
  "nadd_add ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3228
%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u n + dest_nadd ua n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3229
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3230
lemma DEF_nadd_add: "nadd_add =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3231
(%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u n + dest_nadd ua n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3232
  by (import hollight DEF_nadd_add)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3233
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3234
lemma NADD_ADD: "ALL (x::nadd) y::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3235
   dest_nadd (nadd_add x y) = (%n::nat. dest_nadd x n + dest_nadd y n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3236
  by (import hollight NADD_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3237
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3238
lemma NADD_ADD_WELLDEF: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3239
   nadd_eq x x' & nadd_eq y y' --> nadd_eq (nadd_add x y) (nadd_add x' y')"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3240
  by (import hollight NADD_ADD_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3241
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3242
lemma NADD_ADD_SYM: "ALL (x::nadd) y::nadd. nadd_eq (nadd_add x y) (nadd_add y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3243
  by (import hollight NADD_ADD_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3244
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3245
lemma NADD_ADD_ASSOC: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3246
   nadd_eq (nadd_add x (nadd_add y z)) (nadd_add (nadd_add x y) z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3247
  by (import hollight NADD_ADD_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3248
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3249
lemma NADD_ADD_LID: "ALL x::nadd. nadd_eq (nadd_add (nadd_of_num 0) x) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3250
  by (import hollight NADD_ADD_LID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3251
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3252
lemma NADD_ADD_LCANCEL: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3253
   nadd_eq (nadd_add x y) (nadd_add x z) --> nadd_eq y z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3254
  by (import hollight NADD_ADD_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3255
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3256
lemma NADD_LE_ADD: "ALL (x::nadd) y::nadd. nadd_le x (nadd_add x y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3257
  by (import hollight NADD_LE_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3258
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3259
lemma NADD_LE_EXISTS: "ALL (x::nadd) y::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3260
   nadd_le x y --> (EX d::nadd. nadd_eq y (nadd_add x d))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3261
  by (import hollight NADD_LE_EXISTS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3262
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3263
lemma NADD_OF_NUM_ADD: "ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3264
   nadd_eq (nadd_add (nadd_of_num x) (nadd_of_num xa))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3265
    (nadd_of_num (x + xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3266
  by (import hollight NADD_OF_NUM_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3267
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3268
definition nadd_mul :: "nadd => nadd => nadd" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3269
  "nadd_mul ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3270
%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u (dest_nadd ua n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3271
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3272
lemma DEF_nadd_mul: "nadd_mul =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3273
(%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u (dest_nadd ua n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3274
  by (import hollight DEF_nadd_mul)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3275
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3276
lemma NADD_MUL: "ALL (x::nadd) y::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3277
   dest_nadd (nadd_mul x y) = (%n::nat. dest_nadd x (dest_nadd y n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3278
  by (import hollight NADD_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3279
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3280
lemma NADD_MUL_SYM: "ALL (x::nadd) y::nadd. nadd_eq (nadd_mul x y) (nadd_mul y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3281
  by (import hollight NADD_MUL_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3282
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3283
lemma NADD_MUL_ASSOC: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3284
   nadd_eq (nadd_mul x (nadd_mul y z)) (nadd_mul (nadd_mul x y) z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3285
  by (import hollight NADD_MUL_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3286
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3287
lemma NADD_MUL_LID: "ALL x::nadd. nadd_eq (nadd_mul (nadd_of_num (NUMERAL_BIT1 0)) x) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3288
  by (import hollight NADD_MUL_LID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3289
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3290
lemma NADD_LDISTRIB: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3291
   nadd_eq (nadd_mul x (nadd_add y z))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3292
    (nadd_add (nadd_mul x y) (nadd_mul x z))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3293
  by (import hollight NADD_LDISTRIB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3294
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3295
lemma NADD_MUL_WELLDEF_LEMMA: "ALL (x::nadd) (y::nadd) y'::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3296
   nadd_eq y y' --> nadd_eq (nadd_mul x y) (nadd_mul x y')"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3297
  by (import hollight NADD_MUL_WELLDEF_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3298
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3299
lemma NADD_MUL_WELLDEF: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3300
   nadd_eq x x' & nadd_eq y y' --> nadd_eq (nadd_mul x y) (nadd_mul x' y')"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3301
  by (import hollight NADD_MUL_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3302
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3303
lemma NADD_OF_NUM_MUL: "ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3304
   nadd_eq (nadd_mul (nadd_of_num x) (nadd_of_num xa))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3305
    (nadd_of_num (x * xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3306
  by (import hollight NADD_OF_NUM_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3307
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3308
lemma NADD_LE_0: "All (nadd_le (nadd_of_num 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3309
  by (import hollight NADD_LE_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3310
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3311
lemma NADD_EQ_IMP_LE: "ALL (x::nadd) y::nadd. nadd_eq x y --> nadd_le x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3312
  by (import hollight NADD_EQ_IMP_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3313
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3314
lemma NADD_LE_LMUL: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3315
   nadd_le y z --> nadd_le (nadd_mul x y) (nadd_mul x z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3316
  by (import hollight NADD_LE_LMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3317
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3318
lemma NADD_LE_RMUL: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3319
   nadd_le x y --> nadd_le (nadd_mul x z) (nadd_mul y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3320
  by (import hollight NADD_LE_RMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3321
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3322
lemma NADD_LE_RADD: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3323
   nadd_le (nadd_add x z) (nadd_add y z) = nadd_le x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3324
  by (import hollight NADD_LE_RADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3325
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3326
lemma NADD_LE_LADD: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3327
   nadd_le (nadd_add x y) (nadd_add x z) = nadd_le y z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3328
  by (import hollight NADD_LE_LADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3329
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3330
lemma NADD_RDISTRIB: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3331
   nadd_eq (nadd_mul (nadd_add x y) z)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3332
    (nadd_add (nadd_mul x z) (nadd_mul y z))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3333
  by (import hollight NADD_RDISTRIB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3334
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3335
lemma NADD_ARCH_MULT: "ALL (x::nadd) k::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3336
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3337
   (EX xa::nat. nadd_le (nadd_of_num k) (nadd_mul (nadd_of_num xa) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3338
  by (import hollight NADD_ARCH_MULT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3339
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3340
lemma NADD_ARCH_ZERO: "ALL (x::nadd) k::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3341
   (ALL n::nat. nadd_le (nadd_mul (nadd_of_num n) x) k) -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3342
   nadd_eq x (nadd_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3343
  by (import hollight NADD_ARCH_ZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3344
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3345
lemma NADD_ARCH_LEMMA: "ALL (x::nadd) (y::nadd) z::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3346
   (ALL n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3347
       nadd_le (nadd_mul (nadd_of_num n) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3348
        (nadd_add (nadd_mul (nadd_of_num n) y) z)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3349
   nadd_le x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3350
  by (import hollight NADD_ARCH_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3351
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3352
lemma NADD_COMPLETE: "ALL P::nadd => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3353
   Ex P & (EX M::nadd. ALL x::nadd. P x --> nadd_le x M) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3354
   (EX M::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3355
       (ALL x::nadd. P x --> nadd_le x M) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3356
       (ALL M'::nadd. (ALL x::nadd. P x --> nadd_le x M') --> nadd_le M M'))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3357
  by (import hollight NADD_COMPLETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3358
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3359
lemma NADD_UBOUND: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3360
   EX (xa::nat) N::nat. ALL n::nat. <= N n --> <= (dest_nadd x n) (xa * n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3361
  by (import hollight NADD_UBOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3362
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3363
lemma NADD_NONZERO: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3364
   ~ nadd_eq x (nadd_of_num 0) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3365
   (EX N::nat. ALL n::nat. <= N n --> dest_nadd x n ~= 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3366
  by (import hollight NADD_NONZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3367
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3368
lemma NADD_LBOUND: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3369
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3370
   (EX (A::nat) N::nat. ALL n::nat. <= N n --> <= n (A * dest_nadd x n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3371
  by (import hollight NADD_LBOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3372
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3373
definition nadd_rinv :: "nadd => nat => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3374
  "nadd_rinv == %(u::nadd) n::nat. DIV (n * n) (dest_nadd u n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3375
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3376
lemma DEF_nadd_rinv: "nadd_rinv = (%(u::nadd) n::nat. DIV (n * n) (dest_nadd u n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3377
  by (import hollight DEF_nadd_rinv)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3378
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3379
lemma NADD_MUL_LINV_LEMMA0: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3380
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3381
   (EX (xa::nat) B::nat. ALL i::nat. <= (nadd_rinv x i) (xa * i + B))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3382
  by (import hollight NADD_MUL_LINV_LEMMA0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3383
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3384
lemma NADD_MUL_LINV_LEMMA1: "ALL (x::nadd) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3385
   dest_nadd x n ~= 0 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3386
   <= (dist (dest_nadd x n * nadd_rinv x n, n * n)) (dest_nadd x n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3387
  by (import hollight NADD_MUL_LINV_LEMMA1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3388
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3389
lemma NADD_MUL_LINV_LEMMA2: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3390
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3391
   (EX N::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3392
       ALL n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3393
          <= N n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3394
          <= (dist (dest_nadd x n * nadd_rinv x n, n * n)) (dest_nadd x n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3395
  by (import hollight NADD_MUL_LINV_LEMMA2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3396
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3397
lemma NADD_MUL_LINV_LEMMA3: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3398
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3399
   (EX N::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3400
       ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3401
          <= N n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3402
          <= (dist
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3403
               (m * (dest_nadd x m * (dest_nadd x n * nadd_rinv x n)),
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3404
                m * (dest_nadd x m * (n * n))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3405
           (m * (dest_nadd x m * dest_nadd x n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3406
  by (import hollight NADD_MUL_LINV_LEMMA3)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3407
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3408
lemma NADD_MUL_LINV_LEMMA4: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3409
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3410
   (EX N::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3411
       ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3412
          <= N m & <= N n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3413
          <= (dest_nadd x m * dest_nadd x n *
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3414
              dist (m * nadd_rinv x n, n * nadd_rinv x m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3415
           (m * n * dist (m * dest_nadd x n, n * dest_nadd x m) +
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3416
            dest_nadd x m * dest_nadd x n * (m + n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3417
  by (import hollight NADD_MUL_LINV_LEMMA4)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3418
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3419
lemma NADD_MUL_LINV_LEMMA5: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3420
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3421
   (EX (B::nat) N::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3422
       ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3423
          <= N m & <= N n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3424
          <= (dest_nadd x m * dest_nadd x n *
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3425
              dist (m * nadd_rinv x n, n * nadd_rinv x m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3426
           (B * (m * n * (m + n))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3427
  by (import hollight NADD_MUL_LINV_LEMMA5)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3428
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3429
lemma NADD_MUL_LINV_LEMMA6: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3430
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3431
   (EX (B::nat) N::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3432
       ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3433
          <= N m & <= N n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3434
          <= (m * n * dist (m * nadd_rinv x n, n * nadd_rinv x m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3435
           (B * (m * n * (m + n))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3436
  by (import hollight NADD_MUL_LINV_LEMMA6)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3437
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3438
lemma NADD_MUL_LINV_LEMMA7: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3439
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3440
   (EX (B::nat) N::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3441
       ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3442
          <= N m & <= N n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3443
          <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3444
  by (import hollight NADD_MUL_LINV_LEMMA7)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3445
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3446
lemma NADD_MUL_LINV_LEMMA7a: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3447
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3448
   (ALL N::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3449
       EX (A::nat) B::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3450
          ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3451
             <= m N -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3452
             <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (A * n + B))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3453
  by (import hollight NADD_MUL_LINV_LEMMA7a)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3454
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3455
lemma NADD_MUL_LINV_LEMMA8: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3456
   ~ nadd_eq x (nadd_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3457
   (EX B::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3458
       ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3459
          <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3460
  by (import hollight NADD_MUL_LINV_LEMMA8)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3461
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3462
definition nadd_inv :: "nadd => nadd" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3463
  "nadd_inv ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3464
%u::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3465
   COND (nadd_eq u (nadd_of_num 0)) (nadd_of_num 0) (mk_nadd (nadd_rinv u))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3466
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3467
lemma DEF_nadd_inv: "nadd_inv =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3468
(%u::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3469
    COND (nadd_eq u (nadd_of_num 0)) (nadd_of_num 0)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3470
     (mk_nadd (nadd_rinv u)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3471
  by (import hollight DEF_nadd_inv)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3472
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3473
lemma NADD_INV: "ALL x::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3474
   dest_nadd (nadd_inv x) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3475
   COND (nadd_eq x (nadd_of_num 0)) (%n::nat. 0) (nadd_rinv x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3476
  by (import hollight NADD_INV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3477
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3478
lemma NADD_MUL_LINV: "ALL x::nadd.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3479
   ~ nadd_eq x (nadd_of_num 0) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3480
   nadd_eq (nadd_mul (nadd_inv x) x) (nadd_of_num (NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3481
  by (import hollight NADD_MUL_LINV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3482
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3483
lemma NADD_INV_0: "nadd_eq (nadd_inv (nadd_of_num 0)) (nadd_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3484
  by (import hollight NADD_INV_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3485
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3486
lemma NADD_INV_WELLDEF: "ALL (x::nadd) y::nadd. nadd_eq x y --> nadd_eq (nadd_inv x) (nadd_inv y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3487
  by (import hollight NADD_INV_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3488
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3489
typedef (open) hreal = "{s::nadd => bool. EX x::nadd. s = nadd_eq x}"  morphisms "dest_hreal" "mk_hreal"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3490
  apply (rule light_ex_imp_nonempty[where t="nadd_eq (x::nadd)"])
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3491
  by (import hollight TYDEF_hreal)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3492
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3493
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3494
  dest_hreal :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3495
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3496
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3497
  mk_hreal :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3498
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3499
lemmas "TYDEF_hreal_@intern" = typedef_hol2hollight 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3500
  [where a="a :: hreal" and r=r ,
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3501
   OF type_definition_hreal]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3502
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3503
definition hreal_of_num :: "nat => hreal" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3504
  "hreal_of_num == %m::nat. mk_hreal (nadd_eq (nadd_of_num m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3505
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3506
lemma DEF_hreal_of_num: "hreal_of_num = (%m::nat. mk_hreal (nadd_eq (nadd_of_num m)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3507
  by (import hollight DEF_hreal_of_num)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3508
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3509
definition hreal_add :: "hreal => hreal => hreal" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3510
  "hreal_add ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3511
%(x::hreal) y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3512
   mk_hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3513
    (%u::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3514
        EX (xa::nadd) ya::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3515
           nadd_eq (nadd_add xa ya) u & dest_hreal x xa & dest_hreal y ya)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3516
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3517
lemma DEF_hreal_add: "hreal_add =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3518
(%(x::hreal) y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3519
    mk_hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3520
     (%u::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3521
         EX (xa::nadd) ya::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3522
            nadd_eq (nadd_add xa ya) u & dest_hreal x xa & dest_hreal y ya))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3523
  by (import hollight DEF_hreal_add)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3524
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3525
definition hreal_mul :: "hreal => hreal => hreal" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3526
  "hreal_mul ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3527
%(x::hreal) y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3528
   mk_hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3529
    (%u::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3530
        EX (xa::nadd) ya::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3531
           nadd_eq (nadd_mul xa ya) u & dest_hreal x xa & dest_hreal y ya)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3532
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3533
lemma DEF_hreal_mul: "hreal_mul =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3534
(%(x::hreal) y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3535
    mk_hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3536
     (%u::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3537
         EX (xa::nadd) ya::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3538
            nadd_eq (nadd_mul xa ya) u & dest_hreal x xa & dest_hreal y ya))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3539
  by (import hollight DEF_hreal_mul)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3540
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3541
definition hreal_le :: "hreal => hreal => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3542
  "hreal_le ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3543
%(x::hreal) y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3544
   SOME u::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3545
      EX (xa::nadd) ya::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3546
         nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3547
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3548
lemma DEF_hreal_le: "hreal_le =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3549
(%(x::hreal) y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3550
    SOME u::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3551
       EX (xa::nadd) ya::nadd.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3552
          nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3553
  by (import hollight DEF_hreal_le)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3554
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3555
definition hreal_inv :: "hreal => hreal" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3556
  "hreal_inv ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3557
%x::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3558
   mk_hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3559
    (%u::nadd. EX xa::nadd. nadd_eq (nadd_inv xa) u & dest_hreal x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3560
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3561
lemma DEF_hreal_inv: "hreal_inv =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3562
(%x::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3563
    mk_hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3564
     (%u::nadd. EX xa::nadd. nadd_eq (nadd_inv xa) u & dest_hreal x xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3565
  by (import hollight DEF_hreal_inv)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3566
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3567
lemma HREAL_LE_EXISTS_DEF: "ALL (m::hreal) n::hreal. hreal_le m n = (EX d::hreal. n = hreal_add m d)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3568
  by (import hollight HREAL_LE_EXISTS_DEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3569
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3570
lemma HREAL_EQ_ADD_LCANCEL: "ALL (m::hreal) (n::hreal) p::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3571
   (hreal_add m n = hreal_add m p) = (n = p)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3572
  by (import hollight HREAL_EQ_ADD_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3573
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3574
lemma HREAL_EQ_ADD_RCANCEL: "ALL (x::hreal) (xa::hreal) xb::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3575
   (hreal_add x xb = hreal_add xa xb) = (x = xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3576
  by (import hollight HREAL_EQ_ADD_RCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3577
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3578
lemma HREAL_LE_ADD_LCANCEL: "ALL (x::hreal) (xa::hreal) xb::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3579
   hreal_le (hreal_add x xa) (hreal_add x xb) = hreal_le xa xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3580
  by (import hollight HREAL_LE_ADD_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3581
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3582
lemma HREAL_LE_ADD_RCANCEL: "ALL (x::hreal) (xa::hreal) xb::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3583
   hreal_le (hreal_add x xb) (hreal_add xa xb) = hreal_le x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3584
  by (import hollight HREAL_LE_ADD_RCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3585
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3586
lemma HREAL_ADD_RID: "ALL x::hreal. hreal_add x (hreal_of_num 0) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3587
  by (import hollight HREAL_ADD_RID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3588
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3589
lemma HREAL_ADD_RDISTRIB: "ALL (x::hreal) (xa::hreal) xb::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3590
   hreal_mul (hreal_add x xa) xb =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3591
   hreal_add (hreal_mul x xb) (hreal_mul xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3592
  by (import hollight HREAL_ADD_RDISTRIB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3593
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3594
lemma HREAL_MUL_LZERO: "ALL m::hreal. hreal_mul (hreal_of_num 0) m = hreal_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3595
  by (import hollight HREAL_MUL_LZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3596
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3597
lemma HREAL_MUL_RZERO: "ALL x::hreal. hreal_mul x (hreal_of_num 0) = hreal_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3598
  by (import hollight HREAL_MUL_RZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3599
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3600
lemma HREAL_ADD_AC: "hreal_add (m::hreal) (n::hreal) = hreal_add n m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3601
hreal_add (hreal_add m n) (p::hreal) = hreal_add m (hreal_add n p) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3602
hreal_add m (hreal_add n p) = hreal_add n (hreal_add m p)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3603
  by (import hollight HREAL_ADD_AC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3604
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3605
lemma HREAL_LE_ADD2: "ALL (a::hreal) (b::hreal) (c::hreal) d::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3606
   hreal_le a b & hreal_le c d --> hreal_le (hreal_add a c) (hreal_add b d)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3607
  by (import hollight HREAL_LE_ADD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3608
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3609
lemma HREAL_LE_MUL_RCANCEL_IMP: "ALL (a::hreal) (b::hreal) c::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3610
   hreal_le a b --> hreal_le (hreal_mul a c) (hreal_mul b c)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3611
  by (import hollight HREAL_LE_MUL_RCANCEL_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3612
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3613
definition treal_of_num :: "nat => hreal * hreal" where 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3614
  "treal_of_num == %u::nat. (hreal_of_num u, hreal_of_num 0)"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3615
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3616
lemma DEF_treal_of_num: "treal_of_num = (%u::nat. (hreal_of_num u, hreal_of_num 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3617
  by (import hollight DEF_treal_of_num)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3618
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3619
definition treal_neg :: "hreal * hreal => hreal * hreal" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3620
  "treal_neg == %u::hreal * hreal. (snd u, fst u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3621
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3622
lemma DEF_treal_neg: "treal_neg = (%u::hreal * hreal. (snd u, fst u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3623
  by (import hollight DEF_treal_neg)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3624
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3625
definition treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3626
  "treal_add ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3627
%(u::hreal * hreal) ua::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3628
   (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3629
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3630
lemma DEF_treal_add: "treal_add =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3631
(%(u::hreal * hreal) ua::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3632
    (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3633
  by (import hollight DEF_treal_add)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3634
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3635
definition treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3636
  "treal_mul ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3637
%(u::hreal * hreal) ua::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3638
   (hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)),
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3639
    hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3640
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3641
lemma DEF_treal_mul: "treal_mul =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3642
(%(u::hreal * hreal) ua::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3643
    (hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)),
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3644
     hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3645
  by (import hollight DEF_treal_mul)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3646
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3647
definition treal_le :: "hreal * hreal => hreal * hreal => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3648
  "treal_le ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3649
%(u::hreal * hreal) ua::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3650
   hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3651
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3652
lemma DEF_treal_le: "treal_le =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3653
(%(u::hreal * hreal) ua::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3654
    hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3655
  by (import hollight DEF_treal_le)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3656
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3657
definition treal_inv :: "hreal * hreal => hreal * hreal" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3658
  "treal_inv ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3659
%u::hreal * hreal.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3660
   COND (fst u = snd u) (hreal_of_num 0, hreal_of_num 0)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3661
    (COND (hreal_le (snd u) (fst u))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3662
      (hreal_inv (SOME d::hreal. fst u = hreal_add (snd u) d),
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3663
       hreal_of_num 0)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3664
      (hreal_of_num 0,
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3665
       hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3666
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3667
lemma DEF_treal_inv: "treal_inv =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3668
(%u::hreal * hreal.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3669
    COND (fst u = snd u) (hreal_of_num 0, hreal_of_num 0)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3670
     (COND (hreal_le (snd u) (fst u))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3671
       (hreal_inv (SOME d::hreal. fst u = hreal_add (snd u) d),
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3672
        hreal_of_num 0)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3673
       (hreal_of_num 0,
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3674
        hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3675
  by (import hollight DEF_treal_inv)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3676
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3677
definition treal_eq :: "hreal * hreal => hreal * hreal => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3678
  "treal_eq ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3679
%(u::hreal * hreal) ua::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3680
   hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3681
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3682
lemma DEF_treal_eq: "treal_eq =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3683
(%(u::hreal * hreal) ua::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3684
    hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3685
  by (import hollight DEF_treal_eq)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3686
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3687
lemma TREAL_EQ_REFL: "ALL x::hreal * hreal. treal_eq x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3688
  by (import hollight TREAL_EQ_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3689
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3690
lemma TREAL_EQ_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_eq x y = treal_eq y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3691
  by (import hollight TREAL_EQ_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3692
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3693
lemma TREAL_EQ_TRANS: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3694
   treal_eq x y & treal_eq y z --> treal_eq x z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3695
  by (import hollight TREAL_EQ_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3696
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3697
lemma TREAL_EQ_AP: "ALL (x::hreal * hreal) y::hreal * hreal. x = y --> treal_eq x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3698
  by (import hollight TREAL_EQ_AP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3699
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3700
lemma TREAL_OF_NUM_EQ: "ALL (x::nat) xa::nat. treal_eq (treal_of_num x) (treal_of_num xa) = (x = xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3701
  by (import hollight TREAL_OF_NUM_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3702
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3703
lemma TREAL_OF_NUM_LE: "ALL (x::nat) xa::nat. treal_le (treal_of_num x) (treal_of_num xa) = <= x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3704
  by (import hollight TREAL_OF_NUM_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3705
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3706
lemma TREAL_OF_NUM_ADD: "ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3707
   treal_eq (treal_add (treal_of_num x) (treal_of_num xa))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3708
    (treal_of_num (x + xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3709
  by (import hollight TREAL_OF_NUM_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3710
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3711
lemma TREAL_OF_NUM_MUL: "ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3712
   treal_eq (treal_mul (treal_of_num x) (treal_of_num xa))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3713
    (treal_of_num (x * xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3714
  by (import hollight TREAL_OF_NUM_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3715
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3716
lemma TREAL_ADD_SYM_EQ: "ALL (x::hreal * hreal) y::hreal * hreal. treal_add x y = treal_add y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3717
  by (import hollight TREAL_ADD_SYM_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3718
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3719
lemma TREAL_MUL_SYM_EQ: "ALL (x::hreal * hreal) y::hreal * hreal. treal_mul x y = treal_mul y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3720
  by (import hollight TREAL_MUL_SYM_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3721
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3722
lemma TREAL_ADD_SYM: "ALL (x::hreal * hreal) y::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3723
   treal_eq (treal_add x y) (treal_add y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3724
  by (import hollight TREAL_ADD_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3725
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3726
lemma TREAL_ADD_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3727
   treal_eq (treal_add x (treal_add y z)) (treal_add (treal_add x y) z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3728
  by (import hollight TREAL_ADD_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3729
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3730
lemma TREAL_ADD_LID: "ALL x::hreal * hreal. treal_eq (treal_add (treal_of_num 0) x) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3731
  by (import hollight TREAL_ADD_LID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3732
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3733
lemma TREAL_ADD_LINV: "ALL x::hreal * hreal. treal_eq (treal_add (treal_neg x) x) (treal_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3734
  by (import hollight TREAL_ADD_LINV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3735
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3736
lemma TREAL_MUL_SYM: "ALL (x::hreal * hreal) y::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3737
   treal_eq (treal_mul x y) (treal_mul y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3738
  by (import hollight TREAL_MUL_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3739
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3740
lemma TREAL_MUL_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3741
   treal_eq (treal_mul x (treal_mul y z)) (treal_mul (treal_mul x y) z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3742
  by (import hollight TREAL_MUL_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3743
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3744
lemma TREAL_MUL_LID: "ALL x::hreal * hreal.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3745
   treal_eq (treal_mul (treal_of_num (NUMERAL_BIT1 0)) x) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3746
  by (import hollight TREAL_MUL_LID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3747
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3748
lemma TREAL_ADD_LDISTRIB: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3749
   treal_eq (treal_mul x (treal_add y z))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3750
    (treal_add (treal_mul x y) (treal_mul x z))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3751
  by (import hollight TREAL_ADD_LDISTRIB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3752
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3753
lemma TREAL_LE_REFL: "ALL x::hreal * hreal. treal_le x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3754
  by (import hollight TREAL_LE_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3755
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3756
lemma TREAL_LE_ANTISYM: "ALL (x::hreal * hreal) y::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3757
   (treal_le x y & treal_le y x) = treal_eq x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3758
  by (import hollight TREAL_LE_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3759
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3760
lemma TREAL_LE_TRANS: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3761
   treal_le x y & treal_le y z --> treal_le x z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3762
  by (import hollight TREAL_LE_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3763
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3764
lemma TREAL_LE_TOTAL: "ALL (x::hreal * hreal) y::hreal * hreal. treal_le x y | treal_le y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3765
  by (import hollight TREAL_LE_TOTAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3766
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3767
lemma TREAL_LE_LADD_IMP: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3768
   treal_le y z --> treal_le (treal_add x y) (treal_add x z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3769
  by (import hollight TREAL_LE_LADD_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3770
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3771
lemma TREAL_LE_MUL: "ALL (x::hreal * hreal) y::hreal * hreal.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3772
   treal_le (treal_of_num 0) x & treal_le (treal_of_num 0) y -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3773
   treal_le (treal_of_num 0) (treal_mul x y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3774
  by (import hollight TREAL_LE_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3775
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3776
lemma TREAL_INV_0: "treal_eq (treal_inv (treal_of_num 0)) (treal_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3777
  by (import hollight TREAL_INV_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3778
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3779
lemma TREAL_MUL_LINV: "ALL x::hreal * hreal.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3780
   ~ treal_eq x (treal_of_num 0) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3781
   treal_eq (treal_mul (treal_inv x) x) (treal_of_num (NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3782
  by (import hollight TREAL_MUL_LINV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3783
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3784
lemma TREAL_OF_NUM_WELLDEF: "ALL (m::nat) n::nat. m = n --> treal_eq (treal_of_num m) (treal_of_num n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3785
  by (import hollight TREAL_OF_NUM_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3786
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3787
lemma TREAL_NEG_WELLDEF: "ALL (x1::hreal * hreal) x2::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3788
   treal_eq x1 x2 --> treal_eq (treal_neg x1) (treal_neg x2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3789
  by (import hollight TREAL_NEG_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3790
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3791
lemma TREAL_ADD_WELLDEFR: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3792
   treal_eq x1 x2 --> treal_eq (treal_add x1 y) (treal_add x2 y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3793
  by (import hollight TREAL_ADD_WELLDEFR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3794
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3795
lemma TREAL_ADD_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3796
   y2::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3797
   treal_eq x1 x2 & treal_eq y1 y2 -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3798
   treal_eq (treal_add x1 y1) (treal_add x2 y2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3799
  by (import hollight TREAL_ADD_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3800
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3801
lemma TREAL_MUL_WELLDEFR: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3802
   treal_eq x1 x2 --> treal_eq (treal_mul x1 y) (treal_mul x2 y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3803
  by (import hollight TREAL_MUL_WELLDEFR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3804
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3805
lemma TREAL_MUL_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3806
   y2::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3807
   treal_eq x1 x2 & treal_eq y1 y2 -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3808
   treal_eq (treal_mul x1 y1) (treal_mul x2 y2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3809
  by (import hollight TREAL_MUL_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3810
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3811
lemma TREAL_EQ_IMP_LE: "ALL (x::hreal * hreal) y::hreal * hreal. treal_eq x y --> treal_le x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3812
  by (import hollight TREAL_EQ_IMP_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3813
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3814
lemma TREAL_LE_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3815
   y2::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3816
   treal_eq x1 x2 & treal_eq y1 y2 --> treal_le x1 y1 = treal_le x2 y2"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3817
  by (import hollight TREAL_LE_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3818
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3819
lemma TREAL_INV_WELLDEF: "ALL (x::hreal * hreal) y::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3820
   treal_eq x y --> treal_eq (treal_inv x) (treal_inv y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3821
  by (import hollight TREAL_INV_WELLDEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3822
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3823
typedef (open) real = "{s::hreal * hreal => bool. EX x::hreal * hreal. s = treal_eq x}"  morphisms "dest_real" "mk_real"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3824
  apply (rule light_ex_imp_nonempty[where t="treal_eq (x::hreal * hreal)"])
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3825
  by (import hollight TYDEF_real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3826
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3827
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3828
  dest_real :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3829
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3830
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3831
  mk_real :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3832
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3833
lemmas "TYDEF_real_@intern" = typedef_hol2hollight 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3834
  [where a="a :: hollight.real" and r=r ,
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3835
   OF type_definition_real]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3836
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3837
definition real_of_num :: "nat => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3838
  "real_of_num == %m::nat. mk_real (treal_eq (treal_of_num m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3839
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3840
lemma DEF_real_of_num: "real_of_num = (%m::nat. mk_real (treal_eq (treal_of_num m)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3841
  by (import hollight DEF_real_of_num)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3842
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3843
definition real_neg :: "hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3844
  "real_neg ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3845
%x1::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3846
   mk_real
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3847
    (%u::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3848
        EX x1a::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3849
           treal_eq (treal_neg x1a) u & dest_real x1 x1a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3850
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3851
lemma DEF_real_neg: "real_neg =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3852
(%x1::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3853
    mk_real
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3854
     (%u::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3855
         EX x1a::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3856
            treal_eq (treal_neg x1a) u & dest_real x1 x1a))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3857
  by (import hollight DEF_real_neg)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3858
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3859
definition real_add :: "hollight.real => hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3860
  "real_add ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3861
%(x1::hollight.real) y1::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3862
   mk_real
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3863
    (%u::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3864
        EX (x1a::hreal * hreal) y1a::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3865
           treal_eq (treal_add x1a y1a) u &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3866
           dest_real x1 x1a & dest_real y1 y1a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3867
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3868
lemma DEF_real_add: "real_add =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3869
(%(x1::hollight.real) y1::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3870
    mk_real
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3871
     (%u::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3872
         EX (x1a::hreal * hreal) y1a::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3873
            treal_eq (treal_add x1a y1a) u &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3874
            dest_real x1 x1a & dest_real y1 y1a))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3875
  by (import hollight DEF_real_add)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3876
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3877
definition real_mul :: "hollight.real => hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3878
  "real_mul ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3879
%(x1::hollight.real) y1::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3880
   mk_real
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3881
    (%u::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3882
        EX (x1a::hreal * hreal) y1a::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3883
           treal_eq (treal_mul x1a y1a) u &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3884
           dest_real x1 x1a & dest_real y1 y1a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3885
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3886
lemma DEF_real_mul: "real_mul =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3887
(%(x1::hollight.real) y1::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3888
    mk_real
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3889
     (%u::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3890
         EX (x1a::hreal * hreal) y1a::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3891
            treal_eq (treal_mul x1a y1a) u &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3892
            dest_real x1 x1a & dest_real y1 y1a))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3893
  by (import hollight DEF_real_mul)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3894
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3895
definition real_le :: "hollight.real => hollight.real => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3896
  "real_le ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3897
%(x1::hollight.real) y1::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3898
   SOME u::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3899
      EX (x1a::hreal * hreal) y1a::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3900
         treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3901
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3902
lemma DEF_real_le: "real_le =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3903
(%(x1::hollight.real) y1::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3904
    SOME u::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3905
       EX (x1a::hreal * hreal) y1a::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3906
          treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3907
  by (import hollight DEF_real_le)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3908
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3909
definition real_inv :: "hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3910
  "real_inv ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3911
%x::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3912
   mk_real
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3913
    (%u::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3914
        EX xa::hreal * hreal. treal_eq (treal_inv xa) u & dest_real x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3915
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3916
lemma DEF_real_inv: "real_inv =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3917
(%x::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3918
    mk_real
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3919
     (%u::hreal * hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3920
         EX xa::hreal * hreal. treal_eq (treal_inv xa) u & dest_real x xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3921
  by (import hollight DEF_real_inv)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3922
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3923
definition real_sub :: "hollight.real => hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3924
  "real_sub == %(u::hollight.real) ua::hollight.real. real_add u (real_neg ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3925
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3926
lemma DEF_real_sub: "real_sub = (%(u::hollight.real) ua::hollight.real. real_add u (real_neg ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3927
  by (import hollight DEF_real_sub)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3928
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3929
definition real_lt :: "hollight.real => hollight.real => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3930
  "real_lt == %(u::hollight.real) ua::hollight.real. ~ real_le ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3931
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3932
lemma DEF_real_lt: "real_lt = (%(u::hollight.real) ua::hollight.real. ~ real_le ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3933
  by (import hollight DEF_real_lt)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3934
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3935
consts
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3936
  real_ge :: "hollight.real => hollight.real => bool" 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3937
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3938
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3939
  real_ge_def: "hollight.real_ge == %(u::hollight.real) ua::hollight.real. real_le ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3940
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3941
lemma DEF_real_ge: "hollight.real_ge = (%(u::hollight.real) ua::hollight.real. real_le ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3942
  by (import hollight DEF_real_ge)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3943
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3944
consts
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3945
  real_gt :: "hollight.real => hollight.real => bool" 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3946
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3947
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3948
  real_gt_def: "hollight.real_gt == %(u::hollight.real) ua::hollight.real. real_lt ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3949
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3950
lemma DEF_real_gt: "hollight.real_gt = (%(u::hollight.real) ua::hollight.real. real_lt ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3951
  by (import hollight DEF_real_gt)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3952
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3953
definition real_abs :: "hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3954
  "real_abs ==
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3955
%u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3956
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3957
lemma DEF_real_abs: "real_abs =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3958
(%u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3959
  by (import hollight DEF_real_abs)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3960
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3961
definition real_pow :: "hollight.real => nat => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3962
  "real_pow ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3963
SOME real_pow::hollight.real => nat => hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3964
   (ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3965
   (ALL (x::hollight.real) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3966
       real_pow x (Suc n) = real_mul x (real_pow x n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3967
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3968
lemma DEF_real_pow: "real_pow =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3969
(SOME real_pow::hollight.real => nat => hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3970
    (ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3971
    (ALL (x::hollight.real) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3972
        real_pow x (Suc n) = real_mul x (real_pow x n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3973
  by (import hollight DEF_real_pow)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3974
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3975
definition real_div :: "hollight.real => hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3976
  "real_div == %(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3977
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3978
lemma DEF_real_div: "real_div = (%(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3979
  by (import hollight DEF_real_div)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3980
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3981
definition real_max :: "hollight.real => hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3982
  "real_max == %(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3983
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3984
lemma DEF_real_max: "real_max = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3985
  by (import hollight DEF_real_max)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3986
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  3987
definition real_min :: "hollight.real => hollight.real => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3988
  "real_min == %(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3989
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3990
lemma DEF_real_min: "real_min = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3991
  by (import hollight DEF_real_min)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3992
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3993
lemma REAL_HREAL_LEMMA1: "EX x::hreal => hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3994
   (ALL xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3995
       real_le (real_of_num 0) xa = (EX y::hreal. xa = x y)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3996
   (ALL (y::hreal) z::hreal. hreal_le y z = real_le (x y) (x z))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3997
  by (import hollight REAL_HREAL_LEMMA1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3998
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  3999
lemma REAL_HREAL_LEMMA2: "EX (x::hollight.real => hreal) r::hreal => hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4000
   (ALL xa::hreal. x (r xa) = xa) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4001
   (ALL xa::hollight.real. real_le (real_of_num 0) xa --> r (x xa) = xa) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4002
   (ALL x::hreal. real_le (real_of_num 0) (r x)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4003
   (ALL (x::hreal) y::hreal. hreal_le x y = real_le (r x) (r y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4004
  by (import hollight REAL_HREAL_LEMMA2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4005
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4006
lemma REAL_COMPLETE_SOMEPOS: "ALL P::hollight.real => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4007
   (EX x::hollight.real. P x & real_le (real_of_num 0) x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4008
   (EX M::hollight.real. ALL x::hollight.real. P x --> real_le x M) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4009
   (EX M::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4010
       (ALL x::hollight.real. P x --> real_le x M) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4011
       (ALL M'::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4012
           (ALL x::hollight.real. P x --> real_le x M') --> real_le M M'))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4013
  by (import hollight REAL_COMPLETE_SOMEPOS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4014
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4015
lemma REAL_COMPLETE: "ALL P::hollight.real => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4016
   Ex P &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4017
   (EX M::hollight.real. ALL x::hollight.real. P x --> real_le x M) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4018
   (EX M::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4019
       (ALL x::hollight.real. P x --> real_le x M) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4020
       (ALL M'::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4021
           (ALL x::hollight.real. P x --> real_le x M') --> real_le M M'))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4022
  by (import hollight REAL_COMPLETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4023
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4024
lemma REAL_ADD_AC: "real_add (m::hollight.real) (n::hollight.real) = real_add n m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4025
real_add (real_add m n) (p::hollight.real) = real_add m (real_add n p) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4026
real_add m (real_add n p) = real_add n (real_add m p)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4027
  by (import hollight REAL_ADD_AC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4028
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4029
lemma REAL_ADD_RINV: "ALL x::hollight.real. real_add x (real_neg x) = real_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4030
  by (import hollight REAL_ADD_RINV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4031
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4032
lemma REAL_EQ_ADD_LCANCEL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4033
   (real_add x y = real_add x z) = (y = z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4034
  by (import hollight REAL_EQ_ADD_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4035
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4036
lemma REAL_EQ_ADD_RCANCEL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4037
   (real_add x z = real_add y z) = (x = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4038
  by (import hollight REAL_EQ_ADD_RCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4039
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4040
lemma REAL_MUL_RZERO: "ALL x::hollight.real. real_mul x (real_of_num 0) = real_of_num 0"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4041
  by (import hollight REAL_MUL_RZERO)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4042
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4043
lemma REAL_MUL_LZERO: "ALL x::hollight.real. real_mul (real_of_num 0) x = real_of_num 0"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4044
  by (import hollight REAL_MUL_LZERO)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4045
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4046
lemma REAL_NEG_NEG: "ALL x::hollight.real. real_neg (real_neg x) = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4047
  by (import hollight REAL_NEG_NEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4048
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4049
lemma REAL_MUL_RNEG: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4050
   real_mul x (real_neg y) = real_neg (real_mul x y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4051
  by (import hollight REAL_MUL_RNEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4052
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4053
lemma REAL_MUL_LNEG: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4054
   real_mul (real_neg x) y = real_neg (real_mul x y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4055
  by (import hollight REAL_MUL_LNEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4056
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4057
lemma REAL_NEG_ADD: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4058
   real_neg (real_add x y) = real_add (real_neg x) (real_neg y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4059
  by (import hollight REAL_NEG_ADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4060
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4061
lemma REAL_ADD_RID: "ALL x::hollight.real. real_add x (real_of_num 0) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4062
  by (import hollight REAL_ADD_RID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4063
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4064
lemma REAL_NEG_0: "real_neg (real_of_num 0) = real_of_num 0"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4065
  by (import hollight REAL_NEG_0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4066
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4067
lemma REAL_LE_LNEG: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4068
   real_le (real_neg x) y = real_le (real_of_num 0) (real_add x y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4069
  by (import hollight REAL_LE_LNEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4070
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4071
lemma REAL_LE_NEG2: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4072
   real_le (real_neg x) (real_neg y) = real_le y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4073
  by (import hollight REAL_LE_NEG2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4074
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4075
lemma REAL_LE_RNEG: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4076
   real_le x (real_neg y) = real_le (real_add x y) (real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4077
  by (import hollight REAL_LE_RNEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4078
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4079
lemma REAL_OF_NUM_POW: "ALL (x::nat) n::nat. real_pow (real_of_num x) n = real_of_num (EXP x n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4080
  by (import hollight REAL_OF_NUM_POW)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4081
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4082
lemma REAL_POW_NEG: "ALL (x::hollight.real) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4083
   real_pow (real_neg x) n =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4084
   COND (EVEN n) (real_pow x n) (real_neg (real_pow x n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4085
  by (import hollight REAL_POW_NEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4086
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4087
lemma REAL_ABS_NUM: "ALL x::nat. real_abs (real_of_num x) = real_of_num x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4088
  by (import hollight REAL_ABS_NUM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4089
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4090
lemma REAL_ABS_NEG: "ALL x::hollight.real. real_abs (real_neg x) = real_abs x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4091
  by (import hollight REAL_ABS_NEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4092
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4093
lemma REAL_LTE_TOTAL: "ALL (x::hollight.real) xa::hollight.real. real_lt x xa | real_le xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4094
  by (import hollight REAL_LTE_TOTAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4095
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4096
lemma REAL_LET_TOTAL: "ALL (x::hollight.real) xa::hollight.real. real_le x xa | real_lt xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4097
  by (import hollight REAL_LET_TOTAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4098
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4099
lemma REAL_LT_IMP_LE: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> real_le x y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4100
  by (import hollight REAL_LT_IMP_LE)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4101
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4102
lemma REAL_LTE_TRANS: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4103
   real_lt x y & real_le y z --> real_lt x z"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4104
  by (import hollight REAL_LTE_TRANS)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4105
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4106
lemma REAL_LET_TRANS: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4107
   real_le x y & real_lt y z --> real_lt x z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4108
  by (import hollight REAL_LET_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4109
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4110
lemma REAL_LT_TRANS: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4111
   real_lt x y & real_lt y z --> real_lt x z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4112
  by (import hollight REAL_LT_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4113
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4114
lemma REAL_LE_ADD: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4115
   real_le (real_of_num 0) x & real_le (real_of_num 0) y -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4116
   real_le (real_of_num 0) (real_add x y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4117
  by (import hollight REAL_LE_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4118
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4119
lemma REAL_LTE_ANTISYM: "ALL (x::hollight.real) y::hollight.real. ~ (real_lt x y & real_le y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4120
  by (import hollight REAL_LTE_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4121
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4122
lemma REAL_SUB_LE: "ALL (x::hollight.real) xa::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4123
   real_le (real_of_num 0) (real_sub x xa) = real_le xa x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4124
  by (import hollight REAL_SUB_LE)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4125
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4126
lemma REAL_NEG_SUB: "ALL (x::hollight.real) xa::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4127
   real_neg (real_sub x xa) = real_sub xa x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4128
  by (import hollight REAL_NEG_SUB)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4129
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4130
lemma REAL_LE_LT: "ALL (x::hollight.real) xa::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4131
   real_le x xa = (real_lt x xa | x = xa)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4132
  by (import hollight REAL_LE_LT)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4133
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4134
lemma REAL_SUB_LT: "ALL (x::hollight.real) xa::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4135
   real_lt (real_of_num 0) (real_sub x xa) = real_lt xa x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4136
  by (import hollight REAL_SUB_LT)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4137
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4138
lemma REAL_NOT_LT: "ALL (x::hollight.real) xa::hollight.real. (~ real_lt x xa) = real_le xa x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4139
  by (import hollight REAL_NOT_LT)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4140
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4141
lemma REAL_SUB_0: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4142
   (real_sub x y = real_of_num 0) = (x = y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4143
  by (import hollight REAL_SUB_0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4144
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4145
lemma REAL_LT_LE: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4146
   real_lt x y = (real_le x y & x ~= y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4147
  by (import hollight REAL_LT_LE)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4148
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4149
lemma REAL_LT_REFL: "ALL x::hollight.real. ~ real_lt x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4150
  by (import hollight REAL_LT_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4151
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4152
lemma REAL_LTE_ADD: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4153
   real_lt (real_of_num 0) x & real_le (real_of_num 0) y -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4154
   real_lt (real_of_num 0) (real_add x y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4155
  by (import hollight REAL_LTE_ADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4156
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4157
lemma REAL_LET_ADD: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4158
   real_le (real_of_num 0) x & real_lt (real_of_num 0) y -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4159
   real_lt (real_of_num 0) (real_add x y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4160
  by (import hollight REAL_LET_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4161
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4162
lemma REAL_LT_ADD: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4163
   real_lt (real_of_num 0) x & real_lt (real_of_num 0) y -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4164
   real_lt (real_of_num 0) (real_add x y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4165
  by (import hollight REAL_LT_ADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4166
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4167
lemma REAL_ENTIRE: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4168
   (real_mul x y = real_of_num 0) = (x = real_of_num 0 | y = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4169
  by (import hollight REAL_ENTIRE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4170
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4171
lemma REAL_LE_NEGTOTAL: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4172
   real_le (real_of_num 0) x | real_le (real_of_num 0) (real_neg x)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4173
  by (import hollight REAL_LE_NEGTOTAL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4174
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4175
lemma REAL_LE_SQUARE: "ALL x::hollight.real. real_le (real_of_num 0) (real_mul x x)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4176
  by (import hollight REAL_LE_SQUARE)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4177
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4178
lemma REAL_MUL_RID: "ALL x::hollight.real. real_mul x (real_of_num (NUMERAL_BIT1 0)) = x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4179
  by (import hollight REAL_MUL_RID)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4180
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4181
lemma REAL_POW_2: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4182
   real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = real_mul x x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4183
  by (import hollight REAL_POW_2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4184
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4185
lemma REAL_POLY_CLAUSES: "(ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4186
    real_add x (real_add y z) = real_add (real_add x y) z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4187
(ALL (x::hollight.real) y::hollight.real. real_add x y = real_add y x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4188
(ALL x::hollight.real. real_add (real_of_num 0) x = x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4189
(ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4190
    real_mul x (real_mul y z) = real_mul (real_mul x y) z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4191
(ALL (x::hollight.real) y::hollight.real. real_mul x y = real_mul y x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4192
(ALL x::hollight.real. real_mul (real_of_num (NUMERAL_BIT1 0)) x = x) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4193
(ALL x::hollight.real. real_mul (real_of_num 0) x = real_of_num 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4194
(ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4195
    real_mul x (real_add xa xb) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4196
    real_add (real_mul x xa) (real_mul x xb)) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4197
(ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4198
(ALL (x::hollight.real) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4199
    real_pow x (Suc xa) = real_mul x (real_pow x xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4200
  by (import hollight REAL_POLY_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4201
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4202
lemma REAL_POLY_NEG_CLAUSES: "(ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4203
    real_neg x = real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4204
(ALL (x::hollight.real) xa::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4205
    real_sub x xa =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4206
    real_add x (real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) xa))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4207
  by (import hollight REAL_POLY_NEG_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4208
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4209
lemma REAL_POS: "ALL x::nat. real_le (real_of_num 0) (real_of_num x)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4210
  by (import hollight REAL_POS)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4211
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4212
lemma REAL_OF_NUM_LT: "ALL (x::nat) xa::nat. real_lt (real_of_num x) (real_of_num xa) = < x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4213
  by (import hollight REAL_OF_NUM_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4214
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4215
lemma REAL_OF_NUM_GE: "ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4216
   hollight.real_ge (real_of_num x) (real_of_num xa) = >= x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4217
  by (import hollight REAL_OF_NUM_GE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4218
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4219
lemma REAL_OF_NUM_GT: "ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4220
   hollight.real_gt (real_of_num x) (real_of_num xa) = > x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4221
  by (import hollight REAL_OF_NUM_GT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4222
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4223
lemma REAL_OF_NUM_SUC: "ALL x::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4224
   real_add (real_of_num x) (real_of_num (NUMERAL_BIT1 0)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4225
   real_of_num (Suc x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4226
  by (import hollight REAL_OF_NUM_SUC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4227
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4228
lemma REAL_OF_NUM_SUB: "ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4229
   <= m n --> real_sub (real_of_num n) (real_of_num m) = real_of_num (n - m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4230
  by (import hollight REAL_OF_NUM_SUB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4231
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4232
lemma REAL_MUL_AC: "real_mul (m::hollight.real) (n::hollight.real) = real_mul n m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4233
real_mul (real_mul m n) (p::hollight.real) = real_mul m (real_mul n p) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4234
real_mul m (real_mul n p) = real_mul n (real_mul m p)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4235
  by (import hollight REAL_MUL_AC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4236
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4237
lemma REAL_ADD_RDISTRIB: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4238
   real_mul (real_add x y) z = real_add (real_mul x z) (real_mul y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4239
  by (import hollight REAL_ADD_RDISTRIB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4240
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4241
lemma REAL_LT_LADD_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4242
   real_lt y z --> real_lt (real_add x y) (real_add x z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4243
  by (import hollight REAL_LT_LADD_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4244
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4245
lemma REAL_LT_MUL: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4246
   real_lt (real_of_num 0) x & real_lt (real_of_num 0) y -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4247
   real_lt (real_of_num 0) (real_mul x y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4248
  by (import hollight REAL_LT_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4249
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4250
lemma REAL_EQ_ADD_LCANCEL_0: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4251
   (real_add x y = x) = (y = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4252
  by (import hollight REAL_EQ_ADD_LCANCEL_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4253
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4254
lemma REAL_EQ_ADD_RCANCEL_0: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4255
   (real_add x y = y) = (x = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4256
  by (import hollight REAL_EQ_ADD_RCANCEL_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4257
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4258
lemma REAL_LNEG_UNIQ: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4259
   (real_add x y = real_of_num 0) = (x = real_neg y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4260
  by (import hollight REAL_LNEG_UNIQ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4261
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4262
lemma REAL_RNEG_UNIQ: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4263
   (real_add x y = real_of_num 0) = (y = real_neg x)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4264
  by (import hollight REAL_RNEG_UNIQ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4265
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4266
lemma REAL_NEG_LMUL: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4267
   real_neg (real_mul x y) = real_mul (real_neg x) y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4268
  by (import hollight REAL_NEG_LMUL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4269
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4270
lemma REAL_NEG_RMUL: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4271
   real_neg (real_mul x y) = real_mul x (real_neg y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4272
  by (import hollight REAL_NEG_RMUL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4273
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4274
lemma REAL_NEGNEG: "ALL x::hollight.real. real_neg (real_neg x) = x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4275
  by (import hollight REAL_NEGNEG)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4276
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4277
lemma REAL_NEG_MUL2: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4278
   real_mul (real_neg x) (real_neg y) = real_mul x y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4279
  by (import hollight REAL_NEG_MUL2)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4280
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4281
lemma REAL_LT_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4282
   real_lt (real_add x y) (real_add x z) = real_lt y z"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4283
  by (import hollight REAL_LT_LADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4284
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4285
lemma REAL_LT_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4286
   real_lt (real_add x z) (real_add y z) = real_lt x y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4287
  by (import hollight REAL_LT_RADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4288
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4289
lemma REAL_LT_ANTISYM: "ALL (x::hollight.real) y::hollight.real. ~ (real_lt x y & real_lt y x)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4290
  by (import hollight REAL_LT_ANTISYM)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4291
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4292
lemma REAL_LT_GT: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> ~ real_lt y x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4293
  by (import hollight REAL_LT_GT)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4294
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4295
lemma REAL_NOT_EQ: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4296
   (x ~= y) = (real_lt x y | real_lt y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4297
  by (import hollight REAL_NOT_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4298
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4299
lemma REAL_LE_TOTAL: "ALL (x::hollight.real) y::hollight.real. real_le x y | real_le y x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4300
  by (import hollight REAL_LE_TOTAL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4301
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4302
lemma REAL_LE_REFL: "ALL x::hollight.real. real_le x x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4303
  by (import hollight REAL_LE_REFL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4304
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4305
lemma REAL_LE_ANTISYM: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4306
   (real_le x y & real_le y x) = (x = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4307
  by (import hollight REAL_LE_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4308
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4309
lemma REAL_LET_ANTISYM: "ALL (x::hollight.real) y::hollight.real. ~ (real_le x y & real_lt y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4310
  by (import hollight REAL_LET_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4311
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4312
lemma REAL_NEG_LT0: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4313
   real_lt (real_neg x) (real_of_num 0) = real_lt (real_of_num 0) x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4314
  by (import hollight REAL_NEG_LT0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4315
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4316
lemma REAL_NEG_GT0: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4317
   real_lt (real_of_num 0) (real_neg x) = real_lt x (real_of_num 0)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4318
  by (import hollight REAL_NEG_GT0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4319
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4320
lemma REAL_NEG_LE0: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4321
   real_le (real_neg x) (real_of_num 0) = real_le (real_of_num 0) x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4322
  by (import hollight REAL_NEG_LE0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4323
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4324
lemma REAL_NEG_GE0: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4325
   real_le (real_of_num 0) (real_neg x) = real_le x (real_of_num 0)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4326
  by (import hollight REAL_NEG_GE0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4327
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4328
lemma REAL_LT_TOTAL: "ALL (x::hollight.real) y::hollight.real. x = y | real_lt x y | real_lt y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4329
  by (import hollight REAL_LT_TOTAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4330
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4331
lemma REAL_LT_NEGTOTAL: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4332
   x = real_of_num 0 |
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4333
   real_lt (real_of_num 0) x | real_lt (real_of_num 0) (real_neg x)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4334
  by (import hollight REAL_LT_NEGTOTAL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4335
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4336
lemma REAL_LE_01: "real_le (real_of_num 0) (real_of_num (NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4337
  by (import hollight REAL_LE_01)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4338
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4339
lemma REAL_LT_01: "real_lt (real_of_num 0) (real_of_num (NUMERAL_BIT1 0))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4340
  by (import hollight REAL_LT_01)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4341
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4342
lemma REAL_LE_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4343
   real_le (real_add x y) (real_add x z) = real_le y z"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4344
  by (import hollight REAL_LE_LADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4345
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4346
lemma REAL_LE_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4347
   real_le (real_add x z) (real_add y z) = real_le x y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4348
  by (import hollight REAL_LE_RADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4349
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4350
lemma REAL_LT_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4351
   z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4352
   real_lt w x & real_lt y z --> real_lt (real_add w y) (real_add x z)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4353
  by (import hollight REAL_LT_ADD2)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4354
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4355
lemma REAL_LE_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4356
   z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4357
   real_le w x & real_le y z --> real_le (real_add w y) (real_add x z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4358
  by (import hollight REAL_LE_ADD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4359
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4360
lemma REAL_LT_LNEG: "ALL (x::hollight.real) xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4361
   real_lt (real_neg x) xa = real_lt (real_of_num 0) (real_add x xa)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4362
  by (import hollight REAL_LT_LNEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4363
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4364
lemma REAL_LT_RNEG: "ALL (x::hollight.real) xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4365
   real_lt x (real_neg xa) = real_lt (real_add x xa) (real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4366
  by (import hollight REAL_LT_RNEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4367
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4368
lemma REAL_LT_ADDNEG: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4369
   real_lt y (real_add x (real_neg z)) = real_lt (real_add y z) x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4370
  by (import hollight REAL_LT_ADDNEG)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4371
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4372
lemma REAL_LT_ADDNEG2: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4373
   real_lt (real_add x (real_neg y)) z = real_lt x (real_add z y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4374
  by (import hollight REAL_LT_ADDNEG2)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4375
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4376
lemma REAL_LT_ADD1: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4377
   real_le x y --> real_lt x (real_add y (real_of_num (NUMERAL_BIT1 0)))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4378
  by (import hollight REAL_LT_ADD1)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4379
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4380
lemma REAL_SUB_ADD: "ALL (x::hollight.real) y::hollight.real. real_add (real_sub x y) y = x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4381
  by (import hollight REAL_SUB_ADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4382
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4383
lemma REAL_SUB_ADD2: "ALL (x::hollight.real) y::hollight.real. real_add y (real_sub x y) = x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4384
  by (import hollight REAL_SUB_ADD2)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4385
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4386
lemma REAL_SUB_REFL: "ALL x::hollight.real. real_sub x x = real_of_num 0"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4387
  by (import hollight REAL_SUB_REFL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4388
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4389
lemma REAL_LE_DOUBLE: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4390
   real_le (real_of_num 0) (real_add x x) = real_le (real_of_num 0) x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4391
  by (import hollight REAL_LE_DOUBLE)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4392
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4393
lemma REAL_LE_NEGL: "ALL x::hollight.real. real_le (real_neg x) x = real_le (real_of_num 0) x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4394
  by (import hollight REAL_LE_NEGL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4395
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4396
lemma REAL_LE_NEGR: "ALL x::hollight.real. real_le x (real_neg x) = real_le x (real_of_num 0)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4397
  by (import hollight REAL_LE_NEGR)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4398
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4399
lemma REAL_NEG_EQ_0: "ALL x::hollight.real. (real_neg x = real_of_num 0) = (x = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4400
  by (import hollight REAL_NEG_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4401
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4402
lemma REAL_ADD_SUB: "ALL (x::hollight.real) y::hollight.real. real_sub (real_add x y) x = y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4403
  by (import hollight REAL_ADD_SUB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4404
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4405
lemma REAL_NEG_EQ: "ALL (x::hollight.real) y::hollight.real. (real_neg x = y) = (x = real_neg y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4406
  by (import hollight REAL_NEG_EQ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4407
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4408
lemma REAL_NEG_MINUS1: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4409
   real_neg x = real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4410
  by (import hollight REAL_NEG_MINUS1)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4411
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4412
lemma REAL_LT_IMP_NE: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> x ~= y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4413
  by (import hollight REAL_LT_IMP_NE)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4414
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4415
lemma REAL_LE_ADDR: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4416
   real_le x (real_add x y) = real_le (real_of_num 0) y"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4417
  by (import hollight REAL_LE_ADDR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4418
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4419
lemma REAL_LE_ADDL: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4420
   real_le y (real_add x y) = real_le (real_of_num 0) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4421
  by (import hollight REAL_LE_ADDL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4422
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4423
lemma REAL_LT_ADDR: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4424
   real_lt x (real_add x y) = real_lt (real_of_num 0) y"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4425
  by (import hollight REAL_LT_ADDR)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4426
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4427
lemma REAL_LT_ADDL: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4428
   real_lt y (real_add x y) = real_lt (real_of_num 0) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4429
  by (import hollight REAL_LT_ADDL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4430
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4431
lemma REAL_SUB_SUB: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4432
   real_sub (real_sub x y) x = real_neg y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4433
  by (import hollight REAL_SUB_SUB)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4434
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4435
lemma REAL_LT_ADD_SUB: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4436
   real_lt (real_add x y) z = real_lt x (real_sub z y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4437
  by (import hollight REAL_LT_ADD_SUB)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4438
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4439
lemma REAL_LT_SUB_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4440
   real_lt (real_sub x y) z = real_lt x (real_add z y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4441
  by (import hollight REAL_LT_SUB_RADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4442
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4443
lemma REAL_LT_SUB_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4444
   real_lt x (real_sub y z) = real_lt (real_add x z) y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4445
  by (import hollight REAL_LT_SUB_LADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4446
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4447
lemma REAL_LE_SUB_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4448
   real_le x (real_sub y z) = real_le (real_add x z) y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4449
  by (import hollight REAL_LE_SUB_LADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4450
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4451
lemma REAL_LE_SUB_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4452
   real_le (real_sub x y) z = real_le x (real_add z y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4453
  by (import hollight REAL_LE_SUB_RADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4454
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4455
lemma REAL_LT_NEG: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4456
   real_lt (real_neg x) (real_neg y) = real_lt y x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4457
  by (import hollight REAL_LT_NEG)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4458
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4459
lemma REAL_LE_NEG: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4460
   real_le (real_neg x) (real_neg y) = real_le y x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4461
  by (import hollight REAL_LE_NEG)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4462
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4463
lemma REAL_ADD2_SUB2: "ALL (a::hollight.real) (b::hollight.real) (c::hollight.real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4464
   d::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4465
   real_sub (real_add a b) (real_add c d) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4466
   real_add (real_sub a c) (real_sub b d)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4467
  by (import hollight REAL_ADD2_SUB2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4468
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4469
lemma REAL_SUB_LZERO: "ALL x::hollight.real. real_sub (real_of_num 0) x = real_neg x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4470
  by (import hollight REAL_SUB_LZERO)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4471
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4472
lemma REAL_SUB_RZERO: "ALL x::hollight.real. real_sub x (real_of_num 0) = x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4473
  by (import hollight REAL_SUB_RZERO)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4474
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4475
lemma REAL_LET_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4476
   z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4477
   real_le w x & real_lt y z --> real_lt (real_add w y) (real_add x z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4478
  by (import hollight REAL_LET_ADD2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4479
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4480
lemma REAL_LTE_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4481
   z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4482
   real_lt w x & real_le y z --> real_lt (real_add w y) (real_add x z)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4483
  by (import hollight REAL_LTE_ADD2)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4484
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4485
lemma REAL_SUB_LNEG: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4486
   real_sub (real_neg x) y = real_neg (real_add x y)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4487
  by (import hollight REAL_SUB_LNEG)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4488
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4489
lemma REAL_SUB_RNEG: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4490
   real_sub x (real_neg y) = real_add x y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4491
  by (import hollight REAL_SUB_RNEG)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4492
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4493
lemma REAL_SUB_NEG2: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4494
   real_sub (real_neg x) (real_neg y) = real_sub y x"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4495
  by (import hollight REAL_SUB_NEG2)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4496
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4497
lemma REAL_SUB_TRIANGLE: "ALL (a::hollight.real) (b::hollight.real) c::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4498
   real_add (real_sub a b) (real_sub b c) = real_sub a c"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4499
  by (import hollight REAL_SUB_TRIANGLE)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4500
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4501
lemma REAL_EQ_SUB_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4502
   (x = real_sub y z) = (real_add x z = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4503
  by (import hollight REAL_EQ_SUB_LADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4504
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4505
lemma REAL_EQ_SUB_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4506
   (real_sub x y = z) = (x = real_add z y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4507
  by (import hollight REAL_EQ_SUB_RADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4508
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4509
lemma REAL_SUB_SUB2: "ALL (x::hollight.real) y::hollight.real. real_sub x (real_sub x y) = y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4510
  by (import hollight REAL_SUB_SUB2)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4511
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4512
lemma REAL_ADD_SUB2: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4513
   real_sub x (real_add x y) = real_neg y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4514
  by (import hollight REAL_ADD_SUB2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4515
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4516
lemma REAL_EQ_IMP_LE: "ALL (x::hollight.real) y::hollight.real. x = y --> real_le x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4517
  by (import hollight REAL_EQ_IMP_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4518
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4519
lemma REAL_POS_NZ: "ALL x::hollight.real. real_lt (real_of_num 0) x --> x ~= real_of_num 0"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4520
  by (import hollight REAL_POS_NZ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4521
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4522
lemma REAL_DIFFSQ: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4523
   real_mul (real_add x y) (real_sub x y) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4524
   real_sub (real_mul x x) (real_mul y y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4525
  by (import hollight REAL_DIFFSQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4526
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4527
lemma REAL_EQ_NEG2: "ALL (x::hollight.real) y::hollight.real. (real_neg x = real_neg y) = (x = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4528
  by (import hollight REAL_EQ_NEG2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4529
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4530
lemma REAL_LT_NEG2: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4531
   real_lt (real_neg x) (real_neg y) = real_lt y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4532
  by (import hollight REAL_LT_NEG2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4533
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4534
lemma REAL_SUB_LDISTRIB: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4535
   real_mul x (real_sub y z) = real_sub (real_mul x y) (real_mul x z)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4536
  by (import hollight REAL_SUB_LDISTRIB)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4537
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4538
lemma REAL_SUB_RDISTRIB: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4539
   real_mul (real_sub x y) z = real_sub (real_mul x z) (real_mul y z)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4540
  by (import hollight REAL_SUB_RDISTRIB)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4541
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4542
lemma REAL_ABS_ZERO: "ALL x::hollight.real. (real_abs x = real_of_num 0) = (x = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4543
  by (import hollight REAL_ABS_ZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4544
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4545
lemma REAL_ABS_0: "real_abs (real_of_num 0) = real_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4546
  by (import hollight REAL_ABS_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4547
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4548
lemma REAL_ABS_1: "real_abs (real_of_num (NUMERAL_BIT1 0)) = real_of_num (NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4549
  by (import hollight REAL_ABS_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4550
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4551
lemma REAL_ABS_TRIANGLE: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4552
   real_le (real_abs (real_add x y)) (real_add (real_abs x) (real_abs y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4553
  by (import hollight REAL_ABS_TRIANGLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4554
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4555
lemma REAL_ABS_TRIANGLE_LE: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4556
   real_le (real_add (real_abs x) (real_abs (real_sub y x))) z -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4557
   real_le (real_abs y) z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4558
  by (import hollight REAL_ABS_TRIANGLE_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4559
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4560
lemma REAL_ABS_TRIANGLE_LT: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4561
   real_lt (real_add (real_abs x) (real_abs (real_sub y x))) z -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4562
   real_lt (real_abs y) z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4563
  by (import hollight REAL_ABS_TRIANGLE_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4564
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4565
lemma REAL_ABS_POS: "ALL x::hollight.real. real_le (real_of_num 0) (real_abs x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4566
  by (import hollight REAL_ABS_POS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4567
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4568
lemma REAL_ABS_SUB: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4569
   real_abs (real_sub x y) = real_abs (real_sub y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4570
  by (import hollight REAL_ABS_SUB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4571
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4572
lemma REAL_ABS_NZ: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4573
   (x ~= real_of_num 0) = real_lt (real_of_num 0) (real_abs x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4574
  by (import hollight REAL_ABS_NZ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4575
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4576
lemma REAL_ABS_ABS: "ALL x::hollight.real. real_abs (real_abs x) = real_abs x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4577
  by (import hollight REAL_ABS_ABS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4578
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4579
lemma REAL_ABS_LE: "ALL x::hollight.real. real_le x (real_abs x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4580
  by (import hollight REAL_ABS_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4581
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4582
lemma REAL_ABS_REFL: "ALL x::hollight.real. (real_abs x = x) = real_le (real_of_num 0) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4583
  by (import hollight REAL_ABS_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4584
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4585
lemma REAL_ABS_BETWEEN: "ALL (x::hollight.real) (y::hollight.real) d::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4586
   (real_lt (real_of_num 0) d &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4587
    real_lt (real_sub x d) y & real_lt y (real_add x d)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4588
   real_lt (real_abs (real_sub y x)) d"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4589
  by (import hollight REAL_ABS_BETWEEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4590
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4591
lemma REAL_ABS_BOUND: "ALL (x::hollight.real) (y::hollight.real) d::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4592
   real_lt (real_abs (real_sub x y)) d --> real_lt y (real_add x d)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4593
  by (import hollight REAL_ABS_BOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4594
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4595
lemma REAL_ABS_STILLNZ: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4596
   real_lt (real_abs (real_sub x y)) (real_abs y) --> x ~= real_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4597
  by (import hollight REAL_ABS_STILLNZ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4598
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4599
lemma REAL_ABS_CASES: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4600
   x = real_of_num 0 | real_lt (real_of_num 0) (real_abs x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4601
  by (import hollight REAL_ABS_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4602
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4603
lemma REAL_ABS_BETWEEN1: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4604
   real_lt x z & real_lt (real_abs (real_sub y x)) (real_sub z x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4605
   real_lt y z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4606
  by (import hollight REAL_ABS_BETWEEN1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4607
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4608
lemma REAL_ABS_SIGN: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4609
   real_lt (real_abs (real_sub x y)) y --> real_lt (real_of_num 0) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4610
  by (import hollight REAL_ABS_SIGN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4611
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4612
lemma REAL_ABS_SIGN2: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4613
   real_lt (real_abs (real_sub x y)) (real_neg y) -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4614
   real_lt x (real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4615
  by (import hollight REAL_ABS_SIGN2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4616
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4617
lemma REAL_ABS_CIRCLE: "ALL (x::hollight.real) (y::hollight.real) h::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4618
   real_lt (real_abs h) (real_sub (real_abs y) (real_abs x)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4619
   real_lt (real_abs (real_add x h)) (real_abs y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4620
  by (import hollight REAL_ABS_CIRCLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4621
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4622
lemma REAL_SUB_ABS: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4623
   real_le (real_sub (real_abs x) (real_abs y)) (real_abs (real_sub x y))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4624
  by (import hollight REAL_SUB_ABS)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4625
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4626
lemma REAL_ABS_SUB_ABS: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4627
   real_le (real_abs (real_sub (real_abs x) (real_abs y)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4628
    (real_abs (real_sub x y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4629
  by (import hollight REAL_ABS_SUB_ABS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4630
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4631
lemma REAL_ABS_BETWEEN2: "ALL (x0::hollight.real) (x::hollight.real) (y0::hollight.real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4632
   y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4633
   real_lt x0 y0 &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4634
   real_lt
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4635
    (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0)))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4636
      (real_abs (real_sub x x0)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4637
    (real_sub y0 x0) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4638
   real_lt
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4639
    (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0)))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4640
      (real_abs (real_sub y y0)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4641
    (real_sub y0 x0) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4642
   real_lt x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4643
  by (import hollight REAL_ABS_BETWEEN2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4645
lemma REAL_ABS_BOUNDS: "ALL (x::hollight.real) k::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4646
   real_le (real_abs x) k = (real_le (real_neg k) x & real_le x k)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4647
  by (import hollight REAL_ABS_BOUNDS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4648
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4649
lemma REAL_MIN_MAX: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4650
   real_min x y = real_neg (real_max (real_neg x) (real_neg y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4651
  by (import hollight REAL_MIN_MAX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4652
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4653
lemma REAL_MAX_MIN: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4654
   real_max x y = real_neg (real_min (real_neg x) (real_neg y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4655
  by (import hollight REAL_MAX_MIN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4656
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4657
lemma REAL_MAX_MAX: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4658
   real_le x (real_max x y) & real_le y (real_max x y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4659
  by (import hollight REAL_MAX_MAX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4660
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4661
lemma REAL_MIN_MIN: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4662
   real_le (real_min x y) x & real_le (real_min x y) y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4663
  by (import hollight REAL_MIN_MIN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4664
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4665
lemma REAL_MAX_SYM: "ALL (x::hollight.real) y::hollight.real. real_max x y = real_max y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4666
  by (import hollight REAL_MAX_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4667
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4668
lemma REAL_MIN_SYM: "ALL (x::hollight.real) y::hollight.real. real_min x y = real_min y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4669
  by (import hollight REAL_MIN_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4670
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4671
lemma REAL_LE_MAX: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4672
   real_le z (real_max x y) = (real_le z x | real_le z y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4673
  by (import hollight REAL_LE_MAX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4674
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4675
lemma REAL_LE_MIN: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4676
   real_le z (real_min x y) = (real_le z x & real_le z y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4677
  by (import hollight REAL_LE_MIN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4678
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4679
lemma REAL_LT_MAX: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4680
   real_lt z (real_max x y) = (real_lt z x | real_lt z y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4681
  by (import hollight REAL_LT_MAX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4682
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4683
lemma REAL_LT_MIN: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4684
   real_lt z (real_min x y) = (real_lt z x & real_lt z y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4685
  by (import hollight REAL_LT_MIN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4686
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4687
lemma REAL_MAX_LE: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4688
   real_le (real_max x y) z = (real_le x z & real_le y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4689
  by (import hollight REAL_MAX_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4690
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4691
lemma REAL_MIN_LE: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4692
   real_le (real_min x y) z = (real_le x z | real_le y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4693
  by (import hollight REAL_MIN_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4694
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4695
lemma REAL_MAX_LT: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4696
   real_lt (real_max x y) z = (real_lt x z & real_lt y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4697
  by (import hollight REAL_MAX_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4698
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4699
lemma REAL_MIN_LT: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4700
   real_lt (real_min x y) z = (real_lt x z | real_lt y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4701
  by (import hollight REAL_MIN_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4702
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4703
lemma REAL_MAX_ASSOC: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4704
   real_max x (real_max y z) = real_max (real_max x y) z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4705
  by (import hollight REAL_MAX_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4706
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4707
lemma REAL_MIN_ASSOC: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4708
   real_min x (real_min y z) = real_min (real_min x y) z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4709
  by (import hollight REAL_MIN_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4710
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4711
lemma REAL_MAX_ACI: "real_max (x::hollight.real) (y::hollight.real) = real_max y x &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4712
real_max (real_max x y) (z::hollight.real) = real_max x (real_max y z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4713
real_max x (real_max y z) = real_max y (real_max x z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4714
real_max x x = x & real_max x (real_max x y) = real_max x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4715
  by (import hollight REAL_MAX_ACI)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4716
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4717
lemma REAL_MIN_ACI: "real_min (x::hollight.real) (y::hollight.real) = real_min y x &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4718
real_min (real_min x y) (z::hollight.real) = real_min x (real_min y z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4719
real_min x (real_min y z) = real_min y (real_min x z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4720
real_min x x = x & real_min x (real_min x y) = real_min x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4721
  by (import hollight REAL_MIN_ACI)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4722
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4723
lemma REAL_ABS_MUL: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4724
   real_abs (real_mul x y) = real_mul (real_abs x) (real_abs y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4725
  by (import hollight REAL_ABS_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4726
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4727
lemma REAL_POW_LE: "ALL (x::hollight.real) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4728
   real_le (real_of_num 0) x --> real_le (real_of_num 0) (real_pow x n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4729
  by (import hollight REAL_POW_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4730
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4731
lemma REAL_POW_LT: "ALL (x::hollight.real) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4732
   real_lt (real_of_num 0) x --> real_lt (real_of_num 0) (real_pow x n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4733
  by (import hollight REAL_POW_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4734
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4735
lemma REAL_ABS_POW: "ALL (x::hollight.real) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4736
   real_abs (real_pow x n) = real_pow (real_abs x) n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4737
  by (import hollight REAL_ABS_POW)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4738
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4739
lemma REAL_LE_LMUL: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4740
   real_le (real_of_num 0) x & real_le xa xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4741
   real_le (real_mul x xa) (real_mul x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4742
  by (import hollight REAL_LE_LMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4743
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4744
lemma REAL_LE_RMUL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4745
   real_le x y & real_le (real_of_num 0) z -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4746
   real_le (real_mul x z) (real_mul y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4747
  by (import hollight REAL_LE_RMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4748
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4749
lemma REAL_LT_LMUL: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4750
   real_lt (real_of_num 0) x & real_lt xa xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4751
   real_lt (real_mul x xa) (real_mul x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4752
  by (import hollight REAL_LT_LMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4753
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4754
lemma REAL_LT_RMUL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4755
   real_lt x y & real_lt (real_of_num 0) z -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4756
   real_lt (real_mul x z) (real_mul y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4757
  by (import hollight REAL_LT_RMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4758
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4759
lemma REAL_EQ_MUL_LCANCEL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4760
   (real_mul x y = real_mul x z) = (x = real_of_num 0 | y = z)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4761
  by (import hollight REAL_EQ_MUL_LCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4762
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4763
lemma REAL_EQ_MUL_RCANCEL: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4764
   (real_mul x xb = real_mul xa xb) = (x = xa | xb = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4765
  by (import hollight REAL_EQ_MUL_RCANCEL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4766
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4767
lemma REAL_MUL_LINV_UNIQ: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4768
   real_mul x y = real_of_num (NUMERAL_BIT1 0) --> real_inv y = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4769
  by (import hollight REAL_MUL_LINV_UNIQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4770
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4771
lemma REAL_MUL_RINV_UNIQ: "ALL (x::hollight.real) xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4772
   real_mul x xa = real_of_num (NUMERAL_BIT1 0) --> real_inv x = xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4773
  by (import hollight REAL_MUL_RINV_UNIQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4774
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4775
lemma REAL_INV_INV: "ALL x::hollight.real. real_inv (real_inv x) = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4776
  by (import hollight REAL_INV_INV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4777
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4778
lemma REAL_INV_EQ_0: "ALL x::hollight.real. (real_inv x = real_of_num 0) = (x = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4779
  by (import hollight REAL_INV_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4780
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4781
lemma REAL_LT_INV: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4782
   real_lt (real_of_num 0) x --> real_lt (real_of_num 0) (real_inv x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4783
  by (import hollight REAL_LT_INV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4784
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4785
lemma REAL_LT_INV_EQ: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4786
   real_lt (real_of_num 0) (real_inv x) = real_lt (real_of_num 0) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4787
  by (import hollight REAL_LT_INV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4788
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4789
lemma REAL_INV_NEG: "ALL x::hollight.real. real_inv (real_neg x) = real_neg (real_inv x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4790
  by (import hollight REAL_INV_NEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4791
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4792
lemma REAL_LE_INV_EQ: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4793
   real_le (real_of_num 0) (real_inv x) = real_le (real_of_num 0) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4794
  by (import hollight REAL_LE_INV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4795
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4796
lemma REAL_LE_INV: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4797
   real_le (real_of_num 0) x --> real_le (real_of_num 0) (real_inv x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4798
  by (import hollight REAL_LE_INV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4799
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4800
lemma REAL_MUL_RINV: "ALL x::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4801
   x ~= real_of_num 0 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4802
   real_mul x (real_inv x) = real_of_num (NUMERAL_BIT1 0)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4803
  by (import hollight REAL_MUL_RINV)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4804
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4805
lemma REAL_INV_1: "real_inv (real_of_num (NUMERAL_BIT1 0)) = real_of_num (NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4806
  by (import hollight REAL_INV_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4807
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4808
lemma REAL_DIV_1: "ALL x::hollight.real. real_div x (real_of_num (NUMERAL_BIT1 0)) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4809
  by (import hollight REAL_DIV_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4810
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4811
lemma REAL_DIV_REFL: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4812
   x ~= real_of_num 0 --> real_div x x = real_of_num (NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4813
  by (import hollight REAL_DIV_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4814
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4815
lemma REAL_DIV_RMUL: "ALL (x::hollight.real) xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4816
   xa ~= real_of_num 0 --> real_mul (real_div x xa) xa = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4817
  by (import hollight REAL_DIV_RMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4818
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4819
lemma REAL_DIV_LMUL: "ALL (x::hollight.real) xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4820
   xa ~= real_of_num 0 --> real_mul xa (real_div x xa) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4821
  by (import hollight REAL_DIV_LMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4822
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4823
lemma REAL_ABS_INV: "ALL x::hollight.real. real_abs (real_inv x) = real_inv (real_abs x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4824
  by (import hollight REAL_ABS_INV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4825
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4826
lemma REAL_ABS_DIV: "ALL (x::hollight.real) xa::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4827
   real_abs (real_div x xa) = real_div (real_abs x) (real_abs xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4828
  by (import hollight REAL_ABS_DIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4829
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4830
lemma REAL_INV_MUL: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4831
   real_inv (real_mul x y) = real_mul (real_inv x) (real_inv y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4832
  by (import hollight REAL_INV_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4833
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4834
lemma REAL_INV_DIV: "ALL (x::hollight.real) xa::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4835
   real_inv (real_div x xa) = real_div xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4836
  by (import hollight REAL_INV_DIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4837
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4838
lemma REAL_POW_MUL: "ALL (x::hollight.real) (y::hollight.real) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4839
   real_pow (real_mul x y) n = real_mul (real_pow x n) (real_pow y n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4840
  by (import hollight REAL_POW_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4841
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4842
lemma REAL_POW_INV: "ALL (x::hollight.real) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4843
   real_pow (real_inv x) n = real_inv (real_pow x n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4844
  by (import hollight REAL_POW_INV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4845
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4846
lemma REAL_POW_DIV: "ALL (x::hollight.real) (xa::hollight.real) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4847
   real_pow (real_div x xa) xb = real_div (real_pow x xb) (real_pow xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4848
  by (import hollight REAL_POW_DIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4849
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4850
lemma REAL_POW_ADD: "ALL (x::hollight.real) (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4851
   real_pow x (m + n) = real_mul (real_pow x m) (real_pow x n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4852
  by (import hollight REAL_POW_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4853
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4854
lemma REAL_POW_NZ: "ALL (x::hollight.real) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4855
   x ~= real_of_num 0 --> real_pow x n ~= real_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4856
  by (import hollight REAL_POW_NZ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4857
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4858
lemma REAL_POW_SUB: "ALL (x::hollight.real) (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4859
   x ~= real_of_num 0 & <= m n -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4860
   real_pow x (n - m) = real_div (real_pow x n) (real_pow x m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4861
  by (import hollight REAL_POW_SUB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4862
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4863
lemma REAL_LT_IMP_NZ: "ALL x::hollight.real. real_lt (real_of_num 0) x --> x ~= real_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4864
  by (import hollight REAL_LT_IMP_NZ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4865
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4866
lemma REAL_LT_LCANCEL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4867
   real_lt (real_of_num 0) x & real_lt (real_mul x y) (real_mul x z) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4868
   real_lt y z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4869
  by (import hollight REAL_LT_LCANCEL_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4870
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4871
lemma REAL_LT_RCANCEL_IMP: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4872
   real_lt (real_of_num 0) xb & real_lt (real_mul x xb) (real_mul xa xb) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4873
   real_lt x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4874
  by (import hollight REAL_LT_RCANCEL_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4875
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4876
lemma REAL_LE_LCANCEL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4877
   real_lt (real_of_num 0) x & real_le (real_mul x y) (real_mul x z) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4878
   real_le y z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4879
  by (import hollight REAL_LE_LCANCEL_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4880
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4881
lemma REAL_LE_RCANCEL_IMP: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4882
   real_lt (real_of_num 0) xb & real_le (real_mul x xb) (real_mul xa xb) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4883
   real_le x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4884
  by (import hollight REAL_LE_RCANCEL_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4885
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4886
lemma REAL_LE_RMUL_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4887
   real_lt (real_of_num 0) z -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4888
   real_le (real_mul x z) (real_mul y z) = real_le x y"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4889
  by (import hollight REAL_LE_RMUL_EQ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4890
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4891
lemma REAL_LE_LMUL_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4892
   real_lt (real_of_num 0) z -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4893
   real_le (real_mul z x) (real_mul z y) = real_le x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4894
  by (import hollight REAL_LE_LMUL_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4895
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4896
lemma REAL_LT_RMUL_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4897
   real_lt (real_of_num 0) xb -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4898
   real_lt (real_mul x xb) (real_mul xa xb) = real_lt x xa"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4899
  by (import hollight REAL_LT_RMUL_EQ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4900
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4901
lemma REAL_LT_LMUL_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4902
   real_lt (real_of_num 0) xb -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4903
   real_lt (real_mul xb x) (real_mul xb xa) = real_lt x xa"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4904
  by (import hollight REAL_LT_LMUL_EQ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  4905
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4906
lemma REAL_LE_RDIV_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4907
   real_lt (real_of_num 0) z -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4908
   real_le x (real_div y z) = real_le (real_mul x z) y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4909
  by (import hollight REAL_LE_RDIV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4910
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4911
lemma REAL_LE_LDIV_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4912
   real_lt (real_of_num 0) z -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4913
   real_le (real_div x z) y = real_le x (real_mul y z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4914
  by (import hollight REAL_LE_LDIV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4915
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4916
lemma REAL_LT_RDIV_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4917
   real_lt (real_of_num 0) xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4918
   real_lt x (real_div xa xb) = real_lt (real_mul x xb) xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4919
  by (import hollight REAL_LT_RDIV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4920
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4921
lemma REAL_LT_LDIV_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4922
   real_lt (real_of_num 0) xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4923
   real_lt (real_div x xb) xa = real_lt x (real_mul xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4924
  by (import hollight REAL_LT_LDIV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4925
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4926
lemma REAL_EQ_RDIV_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4927
   real_lt (real_of_num 0) xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4928
   (x = real_div xa xb) = (real_mul x xb = xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4929
  by (import hollight REAL_EQ_RDIV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4930
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4931
lemma REAL_EQ_LDIV_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4932
   real_lt (real_of_num 0) xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4933
   (real_div x xb = xa) = (x = real_mul xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4934
  by (import hollight REAL_EQ_LDIV_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4935
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4936
lemma REAL_LT_DIV2_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4937
   real_lt (real_of_num 0) xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4938
   real_lt (real_div x xb) (real_div xa xb) = real_lt x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4939
  by (import hollight REAL_LT_DIV2_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4940
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4941
lemma REAL_LE_DIV2_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4942
   real_lt (real_of_num 0) xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4943
   real_le (real_div x xb) (real_div xa xb) = real_le x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4944
  by (import hollight REAL_LE_DIV2_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4945
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4946
lemma REAL_MUL_2: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4947
   real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) x = real_add x x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4948
  by (import hollight REAL_MUL_2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4949
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4950
lemma REAL_POW_EQ_0: "ALL (x::hollight.real) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4951
   (real_pow x n = real_of_num 0) = (x = real_of_num 0 & n ~= 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4952
  by (import hollight REAL_POW_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4953
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4954
lemma REAL_LE_MUL2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4955
   z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4956
   real_le (real_of_num 0) w &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4957
   real_le w x & real_le (real_of_num 0) y & real_le y z -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4958
   real_le (real_mul w y) (real_mul x z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4959
  by (import hollight REAL_LE_MUL2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4960
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4961
lemma REAL_LT_MUL2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4962
   z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4963
   real_le (real_of_num 0) w &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4964
   real_lt w x & real_le (real_of_num 0) y & real_lt y z -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4965
   real_lt (real_mul w y) (real_mul x z)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4966
  by (import hollight REAL_LT_MUL2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4967
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4968
lemma REAL_LT_SQUARE: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4969
   real_lt (real_of_num 0) (real_mul x x) = (x ~= real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4970
  by (import hollight REAL_LT_SQUARE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4971
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4972
lemma REAL_INV_LE_1: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4973
   real_le (real_of_num (NUMERAL_BIT1 0)) x -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4974
   real_le (real_inv x) (real_of_num (NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4975
  by (import hollight REAL_INV_LE_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4976
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4977
lemma REAL_POW_LE_1: "ALL (n::nat) x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4978
   real_le (real_of_num (NUMERAL_BIT1 0)) x -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4979
   real_le (real_of_num (NUMERAL_BIT1 0)) (real_pow x n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4980
  by (import hollight REAL_POW_LE_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4981
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4982
lemma REAL_POW_1_LE: "ALL (n::nat) x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4983
   real_le (real_of_num 0) x & real_le x (real_of_num (NUMERAL_BIT1 0)) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4984
   real_le (real_pow x n) (real_of_num (NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4985
  by (import hollight REAL_POW_1_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4986
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4987
lemma REAL_POW_1: "ALL x::hollight.real. real_pow x (NUMERAL_BIT1 0) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4988
  by (import hollight REAL_POW_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4989
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4990
lemma REAL_POW_ONE: "ALL n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4991
   real_pow (real_of_num (NUMERAL_BIT1 0)) n = real_of_num (NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4992
  by (import hollight REAL_POW_ONE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4993
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4994
lemma REAL_LT_INV2: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4995
   real_lt (real_of_num 0) x & real_lt x y -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4996
   real_lt (real_inv y) (real_inv x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4997
  by (import hollight REAL_LT_INV2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4998
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  4999
lemma REAL_LE_INV2: "ALL (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5000
   real_lt (real_of_num 0) x & real_le x y -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5001
   real_le (real_inv y) (real_inv x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5002
  by (import hollight REAL_LE_INV2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5003
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5004
lemma REAL_INV_1_LE: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5005
   real_lt (real_of_num 0) x & real_le x (real_of_num (NUMERAL_BIT1 0)) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5006
   real_le (real_of_num (NUMERAL_BIT1 0)) (real_inv x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5007
  by (import hollight REAL_INV_1_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5008
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5009
lemma REAL_SUB_INV: "ALL (x::hollight.real) xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5010
   x ~= real_of_num 0 & xa ~= real_of_num 0 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5011
   real_sub (real_inv x) (real_inv xa) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5012
   real_div (real_sub xa x) (real_mul x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5013
  by (import hollight REAL_SUB_INV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5014
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5015
lemma REAL_DOWN: "ALL d::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5016
   real_lt (real_of_num 0) d -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5017
   (EX x::hollight.real. real_lt (real_of_num 0) x & real_lt x d)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5018
  by (import hollight REAL_DOWN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5019
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5020
lemma REAL_DOWN2: "ALL (d1::hollight.real) d2::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5021
   real_lt (real_of_num 0) d1 & real_lt (real_of_num 0) d2 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5022
   (EX e::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5023
       real_lt (real_of_num 0) e & real_lt e d1 & real_lt e d2)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5024
  by (import hollight REAL_DOWN2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5025
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5026
lemma REAL_POW_LE2: "ALL (n::nat) (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5027
   real_le (real_of_num 0) x & real_le x y -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5028
   real_le (real_pow x n) (real_pow y n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5029
  by (import hollight REAL_POW_LE2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5030
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5031
lemma REAL_POW_MONO: "ALL (m::nat) (n::nat) x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5032
   real_le (real_of_num (NUMERAL_BIT1 0)) x & <= m n -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5033
   real_le (real_pow x m) (real_pow x n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5034
  by (import hollight REAL_POW_MONO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5035
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5036
lemma REAL_POW_LT2: "ALL (n::nat) (x::hollight.real) y::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5037
   n ~= 0 & real_le (real_of_num 0) x & real_lt x y -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5038
   real_lt (real_pow x n) (real_pow y n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5039
  by (import hollight REAL_POW_LT2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5040
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5041
lemma REAL_POW_MONO_LT: "ALL (m::nat) (n::nat) x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5042
   real_lt (real_of_num (NUMERAL_BIT1 0)) x & < m n -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5043
   real_lt (real_pow x m) (real_pow x n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5044
  by (import hollight REAL_POW_MONO_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5045
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5046
lemma REAL_POW_POW: "ALL (x::hollight.real) (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5047
   real_pow (real_pow x m) n = real_pow x (m * n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5048
  by (import hollight REAL_POW_POW)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5049
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5050
lemma REAL_EQ_RCANCEL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5051
   z ~= real_of_num 0 & real_mul x z = real_mul y z --> x = y"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5052
  by (import hollight REAL_EQ_RCANCEL_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5053
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5054
lemma REAL_EQ_LCANCEL_IMP: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5055
   xb ~= real_of_num 0 & real_mul xb x = real_mul xb xa --> x = xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5056
  by (import hollight REAL_EQ_LCANCEL_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5057
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5058
lemma REAL_LT_DIV: "ALL (x::hollight.real) xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5059
   real_lt (real_of_num 0) x & real_lt (real_of_num 0) xa -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5060
   real_lt (real_of_num 0) (real_div x xa)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5061
  by (import hollight REAL_LT_DIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5062
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5063
lemma REAL_LE_DIV: "ALL (x::hollight.real) xa::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5064
   real_le (real_of_num 0) x & real_le (real_of_num 0) xa -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5065
   real_le (real_of_num 0) (real_div x xa)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5066
  by (import hollight REAL_LE_DIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5067
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5068
lemma REAL_DIV_POW2: "ALL (x::hollight.real) (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5069
   x ~= real_of_num 0 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5070
   real_div (real_pow x m) (real_pow x n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5071
   COND (<= n m) (real_pow x (m - n)) (real_inv (real_pow x (n - m)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5072
  by (import hollight REAL_DIV_POW2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5073
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5074
lemma REAL_DIV_POW2_ALT: "ALL (x::hollight.real) (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5075
   x ~= real_of_num 0 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5076
   real_div (real_pow x m) (real_pow x n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5077
   COND (< n m) (real_pow x (m - n)) (real_inv (real_pow x (n - m)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5078
  by (import hollight REAL_DIV_POW2_ALT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5079
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5080
lemma REAL_LT_POW2: "ALL x::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5081
   real_lt (real_of_num 0)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5082
    (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5083
  by (import hollight REAL_LT_POW2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5084
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5085
lemma REAL_LE_POW2: "ALL n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5086
   real_le (real_of_num (NUMERAL_BIT1 0))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5087
    (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5088
  by (import hollight REAL_LE_POW2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5089
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5090
lemma REAL_POW2_ABS: "ALL x::hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5091
   real_pow (real_abs x) (NUMERAL_BIT0 (NUMERAL_BIT1 0)) =
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5092
   real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5093
  by (import hollight REAL_POW2_ABS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5094
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5095
lemma REAL_LE_SQUARE_ABS: "ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5096
   real_le (real_abs x) (real_abs y) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5097
   real_le (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5098
    (real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 0)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5099
  by (import hollight REAL_LE_SQUARE_ABS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5100
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5101
lemma REAL_SOS_EQ_0: "ALL (x::hollight.real) y::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5102
   (real_add (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0)))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5103
     (real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 0))) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5104
    real_of_num 0) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5105
   (x = real_of_num 0 & y = real_of_num 0)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5106
  by (import hollight REAL_SOS_EQ_0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5107
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5108
lemma REAL_WLOG_LE: "(ALL (x::hollight.real) y::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5109
    (P::hollight.real => hollight.real => bool) x y = P y x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5110
(ALL (x::hollight.real) y::hollight.real. real_le x y --> P x y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5111
(ALL x::hollight.real. All (P x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5112
  by (import hollight REAL_WLOG_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5113
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5114
lemma REAL_WLOG_LT: "(ALL x::hollight.real. (P::hollight.real => hollight.real => bool) x x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5115
(ALL (x::hollight.real) y::hollight.real. P x y = P y x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5116
(ALL (x::hollight.real) y::hollight.real. real_lt x y --> P x y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5117
(ALL x::hollight.real. All (P x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5118
  by (import hollight REAL_WLOG_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5119
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5120
definition mod_real :: "hollight.real => hollight.real => hollight.real => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5121
  "mod_real ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5122
%(u::hollight.real) (ua::hollight.real) ub::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5123
   EX q::hollight.real. real_sub ua ub = real_mul q u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5124
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5125
lemma DEF_mod_real: "mod_real =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5126
(%(u::hollight.real) (ua::hollight.real) ub::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5127
    EX q::hollight.real. real_sub ua ub = real_mul q u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5128
  by (import hollight DEF_mod_real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5129
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5130
definition DECIMAL :: "nat => nat => hollight.real" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5131
  "DECIMAL == %(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5132
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5133
lemma DEF_DECIMAL: "DECIMAL = (%(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5134
  by (import hollight DEF_DECIMAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5135
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5136
lemma RAT_LEMMA1: "(y1::hollight.real) ~= real_of_num 0 &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5137
(y2::hollight.real) ~= real_of_num 0 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5138
real_add (real_div (x1::hollight.real) y1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5139
 (real_div (x2::hollight.real) y2) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5140
real_mul (real_add (real_mul x1 y2) (real_mul x2 y1))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5141
 (real_mul (real_inv y1) (real_inv y2))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5142
  by (import hollight RAT_LEMMA1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5143
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5144
lemma RAT_LEMMA2: "real_lt (real_of_num 0) (y1::hollight.real) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5145
real_lt (real_of_num 0) (y2::hollight.real) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5146
real_add (real_div (x1::hollight.real) y1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5147
 (real_div (x2::hollight.real) y2) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5148
real_mul (real_add (real_mul x1 y2) (real_mul x2 y1))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5149
 (real_mul (real_inv y1) (real_inv y2))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5150
  by (import hollight RAT_LEMMA2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5151
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5152
lemma RAT_LEMMA3: "real_lt (real_of_num 0) (y1::hollight.real) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5153
real_lt (real_of_num 0) (y2::hollight.real) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5154
real_sub (real_div (x1::hollight.real) y1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5155
 (real_div (x2::hollight.real) y2) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5156
real_mul (real_sub (real_mul x1 y2) (real_mul x2 y1))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5157
 (real_mul (real_inv y1) (real_inv y2))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5158
  by (import hollight RAT_LEMMA3)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5159
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5160
lemma RAT_LEMMA4: "real_lt (real_of_num 0) (y1::hollight.real) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5161
real_lt (real_of_num 0) (y2::hollight.real) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5162
real_le (real_div (x1::hollight.real) y1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5163
 (real_div (x2::hollight.real) y2) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5164
real_le (real_mul x1 y2) (real_mul x2 y1)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5165
  by (import hollight RAT_LEMMA4)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5166
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5167
lemma RAT_LEMMA5: "real_lt (real_of_num 0) (y1::hollight.real) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5168
real_lt (real_of_num 0) (y2::hollight.real) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5169
(real_div (x1::hollight.real) y1 = real_div (x2::hollight.real) y2) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5170
(real_mul x1 y2 = real_mul x2 y1)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5171
  by (import hollight RAT_LEMMA5)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5172
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5173
definition is_int :: "hollight.real => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5174
  "is_int ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5175
%u::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5176
   EX n::nat. u = real_of_num n | u = real_neg (real_of_num n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5177
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5178
lemma DEF_is_int: "is_int =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5179
(%u::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5180
    EX n::nat. u = real_of_num n | u = real_neg (real_of_num n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5181
  by (import hollight DEF_is_int)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5182
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5183
typedef (open) int = "Collect is_int"  morphisms "dest_int" "mk_int"
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5184
  apply (rule light_ex_imp_nonempty[where t="real_of_num (NUMERAL 0)"])
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5185
  by (import hollight TYDEF_int)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5186
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5187
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5188
  dest_int :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5189
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5190
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5191
  mk_int :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5192
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5193
lemmas "TYDEF_int_@intern" = typedef_hol2hollight 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5194
  [where a="a :: hollight.int" and r=r ,
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5195
   OF type_definition_int]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5196
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5197
lemma dest_int_rep: "ALL x::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5198
   EX n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5199
      dest_int x = real_of_num n | dest_int x = real_neg (real_of_num n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5200
  by (import hollight dest_int_rep)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5201
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5202
definition int_le :: "hollight.int => hollight.int => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5203
  "int_le ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5204
%(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5205
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5206
lemma DEF_int_le: "int_le =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5207
(%(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5208
  by (import hollight DEF_int_le)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5209
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5210
definition int_lt :: "hollight.int => hollight.int => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5211
  "int_lt ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5212
%(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5213
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5214
lemma DEF_int_lt: "int_lt =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5215
(%(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5216
  by (import hollight DEF_int_lt)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5217
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5218
definition int_ge :: "hollight.int => hollight.int => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5219
  "int_ge ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5220
%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5221
   hollight.real_ge (dest_int u) (dest_int ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5222
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5223
lemma DEF_int_ge: "int_ge =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5224
(%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5225
    hollight.real_ge (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5226
  by (import hollight DEF_int_ge)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5227
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5228
definition int_gt :: "hollight.int => hollight.int => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5229
  "int_gt ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5230
%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5231
   hollight.real_gt (dest_int u) (dest_int ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5232
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5233
lemma DEF_int_gt: "int_gt =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5234
(%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5235
    hollight.real_gt (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5236
  by (import hollight DEF_int_gt)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5237
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5238
definition int_of_num :: "nat => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5239
  "int_of_num == %u::nat. mk_int (real_of_num u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5240
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5241
lemma DEF_int_of_num: "int_of_num = (%u::nat. mk_int (real_of_num u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5242
  by (import hollight DEF_int_of_num)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5243
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5244
lemma int_of_num_th: "ALL x::nat. dest_int (int_of_num x) = real_of_num x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5245
  by (import hollight int_of_num_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5246
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5247
definition int_neg :: "hollight.int => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5248
  "int_neg == %u::hollight.int. mk_int (real_neg (dest_int u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5249
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5250
lemma DEF_int_neg: "int_neg = (%u::hollight.int. mk_int (real_neg (dest_int u)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5251
  by (import hollight DEF_int_neg)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5252
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5253
lemma int_neg_th: "ALL x::hollight.int. dest_int (int_neg x) = real_neg (dest_int x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5254
  by (import hollight int_neg_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5255
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5256
definition int_add :: "hollight.int => hollight.int => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5257
  "int_add ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5258
%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5259
   mk_int (real_add (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5260
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5261
lemma DEF_int_add: "int_add =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5262
(%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5263
    mk_int (real_add (dest_int u) (dest_int ua)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5264
  by (import hollight DEF_int_add)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5265
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5266
lemma int_add_th: "ALL (x::hollight.int) xa::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5267
   dest_int (int_add x xa) = real_add (dest_int x) (dest_int xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5268
  by (import hollight int_add_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5269
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5270
definition int_sub :: "hollight.int => hollight.int => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5271
  "int_sub ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5272
%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5273
   mk_int (real_sub (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5274
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5275
lemma DEF_int_sub: "int_sub =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5276
(%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5277
    mk_int (real_sub (dest_int u) (dest_int ua)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5278
  by (import hollight DEF_int_sub)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5279
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5280
lemma int_sub_th: "ALL (x::hollight.int) xa::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5281
   dest_int (int_sub x xa) = real_sub (dest_int x) (dest_int xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5282
  by (import hollight int_sub_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5283
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5284
definition int_mul :: "hollight.int => hollight.int => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5285
  "int_mul ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5286
%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5287
   mk_int (real_mul (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5288
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5289
lemma DEF_int_mul: "int_mul =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5290
(%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5291
    mk_int (real_mul (dest_int u) (dest_int ua)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5292
  by (import hollight DEF_int_mul)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5293
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5294
lemma int_mul_th: "ALL (x::hollight.int) y::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5295
   dest_int (int_mul x y) = real_mul (dest_int x) (dest_int y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5296
  by (import hollight int_mul_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5297
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5298
definition int_abs :: "hollight.int => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5299
  "int_abs == %u::hollight.int. mk_int (real_abs (dest_int u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5300
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5301
lemma DEF_int_abs: "int_abs = (%u::hollight.int. mk_int (real_abs (dest_int u)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5302
  by (import hollight DEF_int_abs)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5303
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5304
lemma int_abs_th: "ALL x::hollight.int. dest_int (int_abs x) = real_abs (dest_int x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5305
  by (import hollight int_abs_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5306
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5307
definition int_max :: "hollight.int => hollight.int => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5308
  "int_max ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5309
%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5310
   mk_int (real_max (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5311
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5312
lemma DEF_int_max: "int_max =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5313
(%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5314
    mk_int (real_max (dest_int u) (dest_int ua)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5315
  by (import hollight DEF_int_max)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5316
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5317
lemma int_max_th: "ALL (x::hollight.int) y::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5318
   dest_int (int_max x y) = real_max (dest_int x) (dest_int y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5319
  by (import hollight int_max_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5320
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5321
definition int_min :: "hollight.int => hollight.int => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5322
  "int_min ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5323
%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5324
   mk_int (real_min (dest_int u) (dest_int ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5325
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5326
lemma DEF_int_min: "int_min =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5327
(%(u::hollight.int) ua::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5328
    mk_int (real_min (dest_int u) (dest_int ua)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5329
  by (import hollight DEF_int_min)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5330
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5331
lemma int_min_th: "ALL (x::hollight.int) y::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5332
   dest_int (int_min x y) = real_min (dest_int x) (dest_int y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5333
  by (import hollight int_min_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5334
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5335
definition int_pow :: "hollight.int => nat => hollight.int" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5336
  "int_pow == %(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5337
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5338
lemma DEF_int_pow: "int_pow = (%(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5339
  by (import hollight DEF_int_pow)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5340
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5341
lemma int_pow_th: "ALL (x::hollight.int) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5342
   dest_int (int_pow x xa) = real_pow (dest_int x) xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5343
  by (import hollight int_pow_th)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5344
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5345
lemma INT_IMAGE: "ALL x::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5346
   (EX n::nat. x = int_of_num n) | (EX n::nat. x = int_neg (int_of_num n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5347
  by (import hollight INT_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5348
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5349
lemma INT_LT_DISCRETE: "ALL (x::hollight.int) y::hollight.int.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5350
   int_lt x y = int_le (int_add x (int_of_num (NUMERAL_BIT1 0))) y"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5351
  by (import hollight INT_LT_DISCRETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5352
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5353
lemma INT_GT_DISCRETE: "ALL (x::hollight.int) xa::hollight.int.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5354
   int_gt x xa = int_ge x (int_add xa (int_of_num (NUMERAL_BIT1 0)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5355
  by (import hollight INT_GT_DISCRETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5356
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5357
lemma INT_FORALL_POS: "(ALL n::nat. (P::hollight.int => bool) (int_of_num n)) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5358
(ALL i::hollight.int. int_le (int_of_num 0) i --> P i)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5359
  by (import hollight INT_FORALL_POS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5360
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5361
lemma INT_ABS_MUL_1: "ALL (x::hollight.int) y::hollight.int.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5362
   (int_abs (int_mul x y) = int_of_num (NUMERAL_BIT1 0)) =
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5363
   (int_abs x = int_of_num (NUMERAL_BIT1 0) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5364
    int_abs y = int_of_num (NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5365
  by (import hollight INT_ABS_MUL_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5366
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5367
lemma INT_POW: "int_pow (x::hollight.int) 0 = int_of_num (NUMERAL_BIT1 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5368
(ALL xa::nat. int_pow x (Suc xa) = int_mul x (int_pow x xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5369
  by (import hollight INT_POW)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5370
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5371
lemma INT_ABS: "ALL x::hollight.int.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5372
   int_abs x = COND (int_le (int_of_num 0) x) x (int_neg x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5373
  by (import hollight INT_ABS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5374
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5375
lemma INT_GE: "ALL (x::hollight.int) xa::hollight.int. int_ge x xa = int_le xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5376
  by (import hollight INT_GE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5377
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5378
lemma INT_GT: "ALL (x::hollight.int) xa::hollight.int. int_gt x xa = int_lt xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5379
  by (import hollight INT_GT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5380
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5381
lemma INT_LT: "ALL (x::hollight.int) xa::hollight.int. int_lt x xa = (~ int_le xa x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5382
  by (import hollight INT_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5383
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5384
lemma INT_ARCH: "ALL (x::hollight.int) d::hollight.int.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5385
   d ~= int_of_num 0 --> (EX c::hollight.int. int_lt x (int_mul c d))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5386
  by (import hollight INT_ARCH)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5387
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5388
definition mod_int :: "hollight.int => hollight.int => hollight.int => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5389
  "mod_int ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5390
%(u::hollight.int) (ua::hollight.int) ub::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5391
   EX q::hollight.int. int_sub ua ub = int_mul q u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5392
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5393
lemma DEF_mod_int: "mod_int =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5394
(%(u::hollight.int) (ua::hollight.int) ub::hollight.int.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5395
    EX q::hollight.int. int_sub ua ub = int_mul q u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5396
  by (import hollight DEF_mod_int)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5397
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5398
definition IN :: "'A => ('A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5399
  "IN == %(u::'A::type) ua::'A::type => bool. ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5400
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5401
lemma DEF_IN: "IN = (%(u::'A::type) ua::'A::type => bool. ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5402
  by (import hollight DEF_IN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5403
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5404
lemma EXTENSION: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5405
   (x = xa) = (ALL xb::'A::type. IN xb x = IN xb xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5406
  by (import hollight EXTENSION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5407
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5408
definition GSPEC :: "('A => bool) => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5409
  "GSPEC == %u::'A::type => bool. u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5410
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5411
lemma DEF_GSPEC: "GSPEC = (%u::'A::type => bool. u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5412
  by (import hollight DEF_GSPEC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5413
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5414
definition SETSPEC :: "'q_37056 => bool => 'q_37056 => bool" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5415
  "SETSPEC == %(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5416
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5417
lemma DEF_SETSPEC: "SETSPEC = (%(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5418
  by (import hollight DEF_SETSPEC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5419
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5420
lemma IN_ELIM_THM: "(ALL (P::(bool => 'q_37089::type => bool) => bool) x::'q_37089::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5421
    IN x (GSPEC (%v::'q_37089::type. P (SETSPEC v))) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5422
    P (%(p::bool) t::'q_37089::type. p & x = t)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5423
(ALL (p::'q_37120::type => bool) x::'q_37120::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5424
    IN x
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5425
     (GSPEC (%v::'q_37120::type. EX y::'q_37120::type. SETSPEC v (p y) y)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5426
    p x) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5427
(ALL (P::(bool => 'q_37148::type => bool) => bool) x::'q_37148::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5428
    GSPEC (%v::'q_37148::type. P (SETSPEC v)) x =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5429
    P (%(p::bool) t::'q_37148::type. p & x = t)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5430
(ALL (p::'q_37177::type => bool) x::'q_37177::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5431
    GSPEC (%v::'q_37177::type. EX y::'q_37177::type. SETSPEC v (p y) y) x =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5432
    p x) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5433
(ALL (p::'q_37194::type => bool) x::'q_37194::type. IN x p = p x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5434
  by (import hollight IN_ELIM_THM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5435
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5436
definition EMPTY :: "'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5437
  "EMPTY == %x::'A::type. False"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5438
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5439
lemma DEF_EMPTY: "EMPTY = (%x::'A::type. False)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5440
  by (import hollight DEF_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5441
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5442
definition INSERT :: "'A => ('A => bool) => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5443
  "INSERT == %(u::'A::type) (ua::'A::type => bool) y::'A::type. IN y ua | y = u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5444
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5445
lemma DEF_INSERT: "INSERT =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5446
(%(u::'A::type) (ua::'A::type => bool) y::'A::type. IN y ua | y = u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5447
  by (import hollight DEF_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5448
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5449
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5450
  UNIV :: "'A => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5451
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5452
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5453
  UNIV_def: "hollight.UNIV == %x::'A::type. True"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5454
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5455
lemma DEF_UNIV: "hollight.UNIV = (%x::'A::type. True)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5456
  by (import hollight DEF_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5457
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5458
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5459
  UNION :: "('A => bool) => ('A => bool) => 'A => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5460
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5461
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5462
  UNION_def: "hollight.UNION ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5463
%(u::'A::type => bool) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5464
   GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u | IN x ua) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5465
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5466
lemma DEF_UNION: "hollight.UNION =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5467
(%(u::'A::type => bool) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5468
    GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u | IN x ua) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5469
  by (import hollight DEF_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5470
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5471
definition UNIONS :: "(('A => bool) => bool) => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5472
  "UNIONS ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5473
%u::('A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5474
   GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5475
    (%ua::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5476
        EX x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5477
           SETSPEC ua (EX ua::'A::type => bool. IN ua u & IN x ua) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5478
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5479
lemma DEF_UNIONS: "UNIONS =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5480
(%u::('A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5481
    GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5482
     (%ua::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5483
         EX x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5484
            SETSPEC ua (EX ua::'A::type => bool. IN ua u & IN x ua) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5485
  by (import hollight DEF_UNIONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5486
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5487
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5488
  INTER :: "('A => bool) => ('A => bool) => 'A => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5489
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5490
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5491
  INTER_def: "hollight.INTER ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5492
%(u::'A::type => bool) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5493
   GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & IN x ua) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5494
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5495
lemma DEF_INTER: "hollight.INTER =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5496
(%(u::'A::type => bool) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5497
    GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & IN x ua) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5498
  by (import hollight DEF_INTER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5499
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5500
definition INTERS :: "(('A => bool) => bool) => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5501
  "INTERS ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5502
%u::('A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5503
   GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5504
    (%ua::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5505
        EX x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5506
           SETSPEC ua (ALL ua::'A::type => bool. IN ua u --> IN x ua) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5507
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5508
lemma DEF_INTERS: "INTERS =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5509
(%u::('A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5510
    GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5511
     (%ua::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5512
         EX x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5513
            SETSPEC ua (ALL ua::'A::type => bool. IN ua u --> IN x ua) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5514
  by (import hollight DEF_INTERS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5515
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5516
definition DIFF :: "('A => bool) => ('A => bool) => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5517
  "DIFF ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5518
%(u::'A::type => bool) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5519
   GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & ~ IN x ua) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5520
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5521
lemma DEF_DIFF: "DIFF =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5522
(%(u::'A::type => bool) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5523
    GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5524
     (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & ~ IN x ua) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5525
  by (import hollight DEF_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5526
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5527
lemma INSERT: "INSERT (x::'A::type) (s::'A::type => bool) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5528
GSPEC (%u::'A::type. EX y::'A::type. SETSPEC u (IN y s | y = x) y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5529
  by (import hollight INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5530
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5531
definition DELETE :: "('A => bool) => 'A => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5532
  "DELETE ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5533
%(u::'A::type => bool) ua::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5534
   GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5535
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5536
lemma DEF_DELETE: "DELETE =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5537
(%(u::'A::type => bool) ua::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5538
    GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5539
  by (import hollight DEF_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5540
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5541
definition SUBSET :: "('A => bool) => ('A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5542
  "SUBSET ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5543
%(u::'A::type => bool) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5544
   ALL x::'A::type. IN x u --> IN x ua"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5545
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5546
lemma DEF_SUBSET: "SUBSET =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5547
(%(u::'A::type => bool) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5548
    ALL x::'A::type. IN x u --> IN x ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5549
  by (import hollight DEF_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5550
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5551
definition PSUBSET :: "('A => bool) => ('A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5552
  "PSUBSET ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5553
%(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5554
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5555
lemma DEF_PSUBSET: "PSUBSET =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5556
(%(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5557
  by (import hollight DEF_PSUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5558
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5559
definition DISJOINT :: "('A => bool) => ('A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5560
  "DISJOINT ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5561
%(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5562
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5563
lemma DEF_DISJOINT: "DISJOINT =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5564
(%(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5565
  by (import hollight DEF_DISJOINT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5566
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5567
definition SING :: "('A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5568
  "SING == %u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5569
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5570
lemma DEF_SING: "SING = (%u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5571
  by (import hollight DEF_SING)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5572
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5573
definition FINITE :: "('A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5574
  "FINITE ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5575
%a::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5576
   ALL FINITE'::('A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5577
      (ALL a::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5578
          a = EMPTY |
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5579
          (EX (x::'A::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5580
              a = INSERT x s & FINITE' s) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5581
          FINITE' a) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5582
      FINITE' a"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5583
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5584
lemma DEF_FINITE: "FINITE =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5585
(%a::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5586
    ALL FINITE'::('A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5587
       (ALL a::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5588
           a = EMPTY |
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5589
           (EX (x::'A::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5590
               a = INSERT x s & FINITE' s) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5591
           FINITE' a) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5592
       FINITE' a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5593
  by (import hollight DEF_FINITE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5594
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5595
definition INFINITE :: "('A => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5596
  "INFINITE == %u::'A::type => bool. ~ FINITE u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5597
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5598
lemma DEF_INFINITE: "INFINITE = (%u::'A::type => bool. ~ FINITE u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5599
  by (import hollight DEF_INFINITE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5600
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5601
definition IMAGE :: "('A => 'B) => ('A => bool) => 'B => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5602
  "IMAGE ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5603
%(u::'A::type => 'B::type) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5604
   GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5605
    (%ub::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5606
        EX y::'B::type. SETSPEC ub (EX x::'A::type. IN x ua & y = u x) y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5607
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5608
lemma DEF_IMAGE: "IMAGE =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5609
(%(u::'A::type => 'B::type) ua::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5610
    GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5611
     (%ub::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5612
         EX y::'B::type. SETSPEC ub (EX x::'A::type. IN x ua & y = u x) y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5613
  by (import hollight DEF_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5614
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5615
definition INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5616
  "INJ ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5617
%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5618
   (ALL x::'A::type. IN x ua --> IN (u x) ub) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5619
   (ALL (x::'A::type) y::'A::type. IN x ua & IN y ua & u x = u y --> x = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5620
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5621
lemma DEF_INJ: "INJ =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5622
(%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5623
    (ALL x::'A::type. IN x ua --> IN (u x) ub) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5624
    (ALL (x::'A::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5625
        IN x ua & IN y ua & u x = u y --> x = y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5626
  by (import hollight DEF_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5627
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5628
definition SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5629
  "SURJ ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5630
%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5631
   (ALL x::'A::type. IN x ua --> IN (u x) ub) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5632
   (ALL x::'B::type. IN x ub --> (EX y::'A::type. IN y ua & u y = x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5633
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5634
lemma DEF_SURJ: "SURJ =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5635
(%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5636
    (ALL x::'A::type. IN x ua --> IN (u x) ub) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5637
    (ALL x::'B::type. IN x ub --> (EX y::'A::type. IN y ua & u y = x)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5638
  by (import hollight DEF_SURJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5639
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5640
definition BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5641
  "BIJ ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5642
%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5643
   INJ u ua ub & SURJ u ua ub"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5645
lemma DEF_BIJ: "BIJ =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5646
(%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5647
    INJ u ua ub & SURJ u ua ub)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5648
  by (import hollight DEF_BIJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5649
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5650
definition CHOICE :: "('A => bool) => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5651
  "CHOICE == %u::'A::type => bool. SOME x::'A::type. IN x u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5652
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5653
lemma DEF_CHOICE: "CHOICE = (%u::'A::type => bool. SOME x::'A::type. IN x u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5654
  by (import hollight DEF_CHOICE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5655
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5656
definition REST :: "('A => bool) => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5657
  "REST == %u::'A::type => bool. DELETE u (CHOICE u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5658
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5659
lemma DEF_REST: "REST = (%u::'A::type => bool. DELETE u (CHOICE u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5660
  by (import hollight DEF_REST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5661
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5662
definition CARD_GE :: "('q_37693 => bool) => ('q_37690 => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5663
  "CARD_GE ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5664
%(u::'q_37693::type => bool) ua::'q_37690::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5665
   EX f::'q_37693::type => 'q_37690::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5666
      ALL y::'q_37690::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5667
         IN y ua --> (EX x::'q_37693::type. IN x u & y = f x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5668
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5669
lemma DEF_CARD_GE: "CARD_GE =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5670
(%(u::'q_37693::type => bool) ua::'q_37690::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5671
    EX f::'q_37693::type => 'q_37690::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5672
       ALL y::'q_37690::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5673
          IN y ua --> (EX x::'q_37693::type. IN x u & y = f x))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5674
  by (import hollight DEF_CARD_GE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5675
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5676
definition CARD_LE :: "('q_37702 => bool) => ('q_37701 => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5677
  "CARD_LE ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5678
%(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5679
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5680
lemma DEF_CARD_LE: "CARD_LE =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5681
(%(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5682
  by (import hollight DEF_CARD_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5683
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5684
definition CARD_EQ :: "('q_37712 => bool) => ('q_37713 => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5685
  "CARD_EQ ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5686
%(u::'q_37712::type => bool) ua::'q_37713::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5687
   CARD_LE u ua & CARD_LE ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5688
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5689
lemma DEF_CARD_EQ: "CARD_EQ =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5690
(%(u::'q_37712::type => bool) ua::'q_37713::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5691
    CARD_LE u ua & CARD_LE ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5692
  by (import hollight DEF_CARD_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5693
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5694
definition CARD_GT :: "('q_37727 => bool) => ('q_37728 => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5695
  "CARD_GT ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5696
%(u::'q_37727::type => bool) ua::'q_37728::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5697
   CARD_GE u ua & ~ CARD_GE ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5698
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5699
lemma DEF_CARD_GT: "CARD_GT =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5700
(%(u::'q_37727::type => bool) ua::'q_37728::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5701
    CARD_GE u ua & ~ CARD_GE ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5702
  by (import hollight DEF_CARD_GT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5703
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5704
definition CARD_LT :: "('q_37743 => bool) => ('q_37744 => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5705
  "CARD_LT ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5706
%(u::'q_37743::type => bool) ua::'q_37744::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5707
   CARD_LE u ua & ~ CARD_LE ua u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5708
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5709
lemma DEF_CARD_LT: "CARD_LT =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5710
(%(u::'q_37743::type => bool) ua::'q_37744::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5711
    CARD_LE u ua & ~ CARD_LE ua u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5712
  by (import hollight DEF_CARD_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5713
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  5714
definition COUNTABLE :: "('q_37757 => bool) => bool" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5715
  "(op ==::(('q_37757::type => bool) => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5716
        => (('q_37757::type => bool) => bool) => prop)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5717
 (COUNTABLE::('q_37757::type => bool) => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5718
 ((CARD_GE::(nat => bool) => ('q_37757::type => bool) => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5719
   (hollight.UNIV::nat => bool))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5720
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5721
lemma DEF_COUNTABLE: "(op =::(('q_37757::type => bool) => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5722
       => (('q_37757::type => bool) => bool) => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5723
 (COUNTABLE::('q_37757::type => bool) => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5724
 ((CARD_GE::(nat => bool) => ('q_37757::type => bool) => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5725
   (hollight.UNIV::nat => bool))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5726
  by (import hollight DEF_COUNTABLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5727
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5728
lemma NOT_IN_EMPTY: "ALL x::'A::type. ~ IN x EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5729
  by (import hollight NOT_IN_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5730
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5731
lemma IN_UNIV: "ALL x::'A::type. IN x hollight.UNIV"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5732
  by (import hollight IN_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5733
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5734
lemma IN_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5735
   IN xb (hollight.UNION x xa) = (IN xb x | IN xb xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5736
  by (import hollight IN_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5737
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5738
lemma IN_UNIONS: "ALL (x::('A::type => bool) => bool) xa::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5739
   IN xa (UNIONS x) = (EX t::'A::type => bool. IN t x & IN xa t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5740
  by (import hollight IN_UNIONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5741
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5742
lemma IN_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5743
   IN xb (hollight.INTER x xa) = (IN xb x & IN xb xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5744
  by (import hollight IN_INTER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5745
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5746
lemma IN_INTERS: "ALL (x::('A::type => bool) => bool) xa::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5747
   IN xa (INTERS x) = (ALL t::'A::type => bool. IN t x --> IN xa t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5748
  by (import hollight IN_INTERS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5749
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5750
lemma IN_DIFF: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5751
   IN xb (DIFF x xa) = (IN xb x & ~ IN xb xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5752
  by (import hollight IN_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5753
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5754
lemma IN_INSERT: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5755
   IN x (INSERT xa xb) = (x = xa | IN x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5756
  by (import hollight IN_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5757
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5758
lemma IN_DELETE: "ALL (x::'A::type => bool) (xa::'A::type) xb::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5759
   IN xa (DELETE x xb) = (IN xa x & xa ~= xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5760
  by (import hollight IN_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5761
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5762
lemma IN_SING: "ALL (x::'A::type) xa::'A::type. IN x (INSERT xa EMPTY) = (x = xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5763
  by (import hollight IN_SING)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5764
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5765
lemma IN_IMAGE: "ALL (x::'B::type) (xa::'A::type => bool) xb::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5766
   IN x (IMAGE xb xa) = (EX xc::'A::type. x = xb xc & IN xc xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5767
  by (import hollight IN_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5768
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5769
lemma IN_REST: "ALL (x::'A::type) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5770
   IN x (REST xa) = (IN x xa & x ~= CHOICE xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5771
  by (import hollight IN_REST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5772
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5773
lemma CHOICE_DEF: "ALL x::'A::type => bool. x ~= EMPTY --> IN (CHOICE x) x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5774
  by (import hollight CHOICE_DEF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5775
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5776
lemma NOT_EQUAL_SETS: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5777
   (x ~= xa) = (EX xb::'A::type. IN xb xa = (~ IN xb x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5778
  by (import hollight NOT_EQUAL_SETS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5779
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5780
lemma MEMBER_NOT_EMPTY: "ALL x::'A::type => bool. (EX xa::'A::type. IN xa x) = (x ~= EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5781
  by (import hollight MEMBER_NOT_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5782
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5783
lemma UNIV_NOT_EMPTY: "(Not::bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5784
 ((op =::('A::type => bool) => ('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5785
   (hollight.UNIV::'A::type => bool) (EMPTY::'A::type => bool))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5786
  by (import hollight UNIV_NOT_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5787
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5788
lemma EMPTY_NOT_UNIV: "(Not::bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5789
 ((op =::('A::type => bool) => ('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5790
   (EMPTY::'A::type => bool) (hollight.UNIV::'A::type => bool))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5791
  by (import hollight EMPTY_NOT_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5792
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5793
lemma EQ_UNIV: "(ALL x::'A::type. IN x (s::'A::type => bool)) = (s = hollight.UNIV)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5794
  by (import hollight EQ_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5795
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5796
lemma SUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5797
   SUBSET x xa & SUBSET xa xb --> SUBSET x xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5798
  by (import hollight SUBSET_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5799
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5800
lemma SUBSET_REFL: "ALL x::'A::type => bool. SUBSET x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5801
  by (import hollight SUBSET_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5802
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5803
lemma SUBSET_ANTISYM: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5804
   SUBSET x xa & SUBSET xa x --> x = xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5805
  by (import hollight SUBSET_ANTISYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5806
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5807
lemma EMPTY_SUBSET: "(All::(('A::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5808
 ((SUBSET::('A::type => bool) => ('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5809
   (EMPTY::'A::type => bool))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5810
  by (import hollight EMPTY_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5811
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5812
lemma SUBSET_EMPTY: "ALL x::'A::type => bool. SUBSET x EMPTY = (x = EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5813
  by (import hollight SUBSET_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5814
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5815
lemma SUBSET_UNIV: "ALL x::'A::type => bool. SUBSET x hollight.UNIV"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5816
  by (import hollight SUBSET_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5817
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5818
lemma UNIV_SUBSET: "ALL x::'A::type => bool. SUBSET hollight.UNIV x = (x = hollight.UNIV)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5819
  by (import hollight UNIV_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5820
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5821
lemma PSUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5822
   PSUBSET x xa & PSUBSET xa xb --> PSUBSET x xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5823
  by (import hollight PSUBSET_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5824
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5825
lemma PSUBSET_SUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5826
   PSUBSET x xa & SUBSET xa xb --> PSUBSET x xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5827
  by (import hollight PSUBSET_SUBSET_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5828
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5829
lemma SUBSET_PSUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5830
   SUBSET x xa & PSUBSET xa xb --> PSUBSET x xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5831
  by (import hollight SUBSET_PSUBSET_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5832
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5833
lemma PSUBSET_IRREFL: "ALL x::'A::type => bool. ~ PSUBSET x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5834
  by (import hollight PSUBSET_IRREFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5835
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5836
lemma NOT_PSUBSET_EMPTY: "ALL x::'A::type => bool. ~ PSUBSET x EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5837
  by (import hollight NOT_PSUBSET_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5838
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5839
lemma NOT_UNIV_PSUBSET: "ALL x::'A::type => bool. ~ PSUBSET hollight.UNIV x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5840
  by (import hollight NOT_UNIV_PSUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5841
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5842
lemma PSUBSET_UNIV: "ALL x::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5843
   PSUBSET x hollight.UNIV = (EX xa::'A::type. ~ IN xa x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5844
  by (import hollight PSUBSET_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5845
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5846
lemma UNION_ASSOC: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5847
   hollight.UNION (hollight.UNION x xa) xb =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5848
   hollight.UNION x (hollight.UNION xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5849
  by (import hollight UNION_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5850
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5851
lemma UNION_IDEMPOT: "ALL x::'A::type => bool. hollight.UNION x x = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5852
  by (import hollight UNION_IDEMPOT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5853
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5854
lemma UNION_COMM: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5855
   hollight.UNION x xa = hollight.UNION xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5856
  by (import hollight UNION_COMM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5857
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5858
lemma SUBSET_UNION: "(ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5859
    SUBSET x (hollight.UNION x xa)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5860
(ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5861
    SUBSET x (hollight.UNION xa x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5862
  by (import hollight SUBSET_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5863
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5864
lemma SUBSET_UNION_ABSORPTION: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5865
   SUBSET x xa = (hollight.UNION x xa = xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5866
  by (import hollight SUBSET_UNION_ABSORPTION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5867
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5868
lemma UNION_EMPTY: "(ALL x::'A::type => bool. hollight.UNION EMPTY x = x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5869
(ALL x::'A::type => bool. hollight.UNION x EMPTY = x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5870
  by (import hollight UNION_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5871
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5872
lemma UNION_UNIV: "(ALL x::'A::type => bool. hollight.UNION hollight.UNIV x = hollight.UNIV) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5873
(ALL x::'A::type => bool. hollight.UNION x hollight.UNIV = hollight.UNIV)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5874
  by (import hollight UNION_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5875
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5876
lemma EMPTY_UNION: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5877
   (hollight.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5878
  by (import hollight EMPTY_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5879
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5880
lemma UNION_SUBSET: "ALL (x::'q_38594::type => bool) (xa::'q_38594::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5881
   xb::'q_38594::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5882
   SUBSET (hollight.UNION x xa) xb = (SUBSET x xb & SUBSET xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5883
  by (import hollight UNION_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5884
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5885
lemma INTER_ASSOC: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5886
   hollight.INTER (hollight.INTER x xa) xb =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5887
   hollight.INTER x (hollight.INTER xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5888
  by (import hollight INTER_ASSOC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5889
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5890
lemma INTER_IDEMPOT: "ALL x::'A::type => bool. hollight.INTER x x = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5891
  by (import hollight INTER_IDEMPOT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5892
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5893
lemma INTER_COMM: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5894
   hollight.INTER x xa = hollight.INTER xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5895
  by (import hollight INTER_COMM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5896
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5897
lemma INTER_SUBSET: "(ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5898
    SUBSET (hollight.INTER x xa) x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5899
(ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5900
    SUBSET (hollight.INTER xa x) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5901
  by (import hollight INTER_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5902
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5903
lemma SUBSET_INTER_ABSORPTION: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5904
   SUBSET x xa = (hollight.INTER x xa = x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5905
  by (import hollight SUBSET_INTER_ABSORPTION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5906
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5907
lemma INTER_EMPTY: "(ALL x::'A::type => bool. hollight.INTER EMPTY x = EMPTY) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5908
(ALL x::'A::type => bool. hollight.INTER x EMPTY = EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5909
  by (import hollight INTER_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5910
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5911
lemma INTER_UNIV: "(ALL x::'A::type => bool. hollight.INTER hollight.UNIV x = x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5912
(ALL x::'A::type => bool. hollight.INTER x hollight.UNIV = x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5913
  by (import hollight INTER_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5914
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5915
lemma UNION_OVER_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5916
   hollight.INTER x (hollight.UNION xa xb) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5917
   hollight.UNION (hollight.INTER x xa) (hollight.INTER x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5918
  by (import hollight UNION_OVER_INTER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5919
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5920
lemma INTER_OVER_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5921
   hollight.UNION x (hollight.INTER xa xb) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5922
   hollight.INTER (hollight.UNION x xa) (hollight.UNION x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5923
  by (import hollight INTER_OVER_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5924
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5925
lemma IN_DISJOINT: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5926
   DISJOINT x xa = (~ (EX xb::'A::type. IN xb x & IN xb xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5927
  by (import hollight IN_DISJOINT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5928
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5929
lemma DISJOINT_SYM: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5930
   DISJOINT x xa = DISJOINT xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5931
  by (import hollight DISJOINT_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5932
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5933
lemma DISJOINT_EMPTY: "ALL x::'A::type => bool. DISJOINT EMPTY x & DISJOINT x EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5934
  by (import hollight DISJOINT_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5935
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5936
lemma DISJOINT_EMPTY_REFL: "ALL x::'A::type => bool. (x = EMPTY) = DISJOINT x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5937
  by (import hollight DISJOINT_EMPTY_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5938
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5939
lemma DISJOINT_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5940
   DISJOINT (hollight.UNION x xa) xb = (DISJOINT x xb & DISJOINT xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5941
  by (import hollight DISJOINT_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5942
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5943
lemma DIFF_EMPTY: "ALL x::'A::type => bool. DIFF x EMPTY = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5944
  by (import hollight DIFF_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5945
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5946
lemma EMPTY_DIFF: "ALL x::'A::type => bool. DIFF EMPTY x = EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5947
  by (import hollight EMPTY_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5948
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5949
lemma DIFF_UNIV: "ALL x::'A::type => bool. DIFF x hollight.UNIV = EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5950
  by (import hollight DIFF_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5951
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5952
lemma DIFF_DIFF: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5953
   DIFF (DIFF x xa) xa = DIFF x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5954
  by (import hollight DIFF_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5955
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5956
lemma DIFF_EQ_EMPTY: "ALL x::'A::type => bool. DIFF x x = EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5957
  by (import hollight DIFF_EQ_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5958
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  5959
lemma SUBSET_DIFF: "ALL (x::'q_39012::type => bool) xa::'q_39012::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5960
   SUBSET (DIFF x xa) x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5961
  by (import hollight SUBSET_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5962
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5963
lemma COMPONENT: "ALL (x::'A::type) s::'A::type => bool. IN x (INSERT x s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5964
  by (import hollight COMPONENT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5965
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5966
lemma DECOMPOSITION: "ALL (s::'A::type => bool) x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5967
   IN x s = (EX t::'A::type => bool. s = INSERT x t & ~ IN x t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5968
  by (import hollight DECOMPOSITION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5969
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5970
lemma SET_CASES: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5971
   s = EMPTY |
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5972
   (EX (x::'A::type) t::'A::type => bool. s = INSERT x t & ~ IN x t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5973
  by (import hollight SET_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5974
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5975
lemma ABSORPTION: "ALL (x::'A::type) xa::'A::type => bool. IN x xa = (INSERT x xa = xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5976
  by (import hollight ABSORPTION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5977
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5978
lemma INSERT_INSERT: "ALL (x::'A::type) xa::'A::type => bool. INSERT x (INSERT x xa) = INSERT x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5979
  by (import hollight INSERT_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5980
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5981
lemma INSERT_COMM: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5982
   INSERT x (INSERT xa xb) = INSERT xa (INSERT x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5983
  by (import hollight INSERT_COMM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5984
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5985
lemma INSERT_UNIV: "ALL x::'A::type. INSERT x hollight.UNIV = hollight.UNIV"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5986
  by (import hollight INSERT_UNIV)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5987
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5988
lemma NOT_INSERT_EMPTY: "ALL (x::'A::type) xa::'A::type => bool. INSERT x xa ~= EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5989
  by (import hollight NOT_INSERT_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5990
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5991
lemma NOT_EMPTY_INSERT: "ALL (x::'A::type) xa::'A::type => bool. EMPTY ~= INSERT x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5992
  by (import hollight NOT_EMPTY_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5993
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5994
lemma INSERT_UNION: "ALL (x::'A::type) (s::'A::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5995
   hollight.UNION (INSERT x s) t =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5996
   COND (IN x t) (hollight.UNION s t) (INSERT x (hollight.UNION s t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5997
  by (import hollight INSERT_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5998
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  5999
lemma INSERT_UNION_EQ: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6000
   hollight.UNION (INSERT x xa) xb = INSERT x (hollight.UNION xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6001
  by (import hollight INSERT_UNION_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6002
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6003
lemma INSERT_INTER: "ALL (x::'A::type) (s::'A::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6004
   hollight.INTER (INSERT x s) t =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6005
   COND (IN x t) (INSERT x (hollight.INTER s t)) (hollight.INTER s t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6006
  by (import hollight INSERT_INTER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6007
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6008
lemma DISJOINT_INSERT: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6009
   DISJOINT (INSERT x xa) xb = (DISJOINT xa xb & ~ IN x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6010
  by (import hollight DISJOINT_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6011
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6012
lemma INSERT_SUBSET: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6013
   SUBSET (INSERT x xa) xb = (IN x xb & SUBSET xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6014
  by (import hollight INSERT_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6015
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6016
lemma SUBSET_INSERT: "ALL (x::'A::type) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6017
   ~ IN x xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6018
   (ALL xb::'A::type => bool. SUBSET xa (INSERT x xb) = SUBSET xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6019
  by (import hollight SUBSET_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6020
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6021
lemma INSERT_DIFF: "ALL (s::'A::type => bool) (t::'A::type => bool) x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6022
   DIFF (INSERT x s) t = COND (IN x t) (DIFF s t) (INSERT x (DIFF s t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6023
  by (import hollight INSERT_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6024
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6025
lemma INSERT_AC: "INSERT (x::'q_39468::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6026
 (INSERT (y::'q_39468::type) (s::'q_39468::type => bool)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6027
INSERT y (INSERT x s) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6028
INSERT x (INSERT x s) = INSERT x s"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6029
  by (import hollight INSERT_AC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6030
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6031
lemma INTER_ACI: "hollight.INTER (p::'q_39535::type => bool) (q::'q_39535::type => bool) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6032
hollight.INTER q p &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6033
hollight.INTER (hollight.INTER p q) (r::'q_39535::type => bool) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6034
hollight.INTER p (hollight.INTER q r) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6035
hollight.INTER p (hollight.INTER q r) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6036
hollight.INTER q (hollight.INTER p r) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6037
hollight.INTER p p = p &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6038
hollight.INTER p (hollight.INTER p q) = hollight.INTER p q"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6039
  by (import hollight INTER_ACI)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6040
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6041
lemma UNION_ACI: "hollight.UNION (p::'q_39601::type => bool) (q::'q_39601::type => bool) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6042
hollight.UNION q p &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6043
hollight.UNION (hollight.UNION p q) (r::'q_39601::type => bool) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6044
hollight.UNION p (hollight.UNION q r) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6045
hollight.UNION p (hollight.UNION q r) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6046
hollight.UNION q (hollight.UNION p r) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6047
hollight.UNION p p = p &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6048
hollight.UNION p (hollight.UNION p q) = hollight.UNION p q"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6049
  by (import hollight UNION_ACI)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6050
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6051
lemma DELETE_NON_ELEMENT: "ALL (x::'A::type) xa::'A::type => bool. (~ IN x xa) = (DELETE xa x = xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6052
  by (import hollight DELETE_NON_ELEMENT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6053
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6054
lemma IN_DELETE_EQ: "ALL (s::'A::type => bool) (x::'A::type) x'::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6055
   (IN x s = IN x' s) = (IN x (DELETE s x') = IN x' (DELETE s x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6056
  by (import hollight IN_DELETE_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6057
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6058
lemma EMPTY_DELETE: "ALL x::'A::type. DELETE EMPTY x = EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6059
  by (import hollight EMPTY_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6060
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6061
lemma DELETE_DELETE: "ALL (x::'A::type) xa::'A::type => bool. DELETE (DELETE xa x) x = DELETE xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6062
  by (import hollight DELETE_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6063
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6064
lemma DELETE_COMM: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6065
   DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6066
  by (import hollight DELETE_COMM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6067
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6068
lemma DELETE_SUBSET: "ALL (x::'A::type) xa::'A::type => bool. SUBSET (DELETE xa x) xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6069
  by (import hollight DELETE_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6070
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6071
lemma SUBSET_DELETE: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6072
   SUBSET xa (DELETE xb x) = (~ IN x xa & SUBSET xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6073
  by (import hollight SUBSET_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6074
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6075
lemma SUBSET_INSERT_DELETE: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6076
   SUBSET xa (INSERT x xb) = SUBSET (DELETE xa x) xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6077
  by (import hollight SUBSET_INSERT_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6078
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6079
lemma DIFF_INSERT: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6080
   DIFF x (INSERT xb xa) = DIFF (DELETE x xb) xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6081
  by (import hollight DIFF_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6082
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6083
lemma PSUBSET_INSERT_SUBSET: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6084
   PSUBSET x xa = (EX xb::'A::type. ~ IN xb x & SUBSET (INSERT xb x) xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6085
  by (import hollight PSUBSET_INSERT_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6086
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6087
lemma PSUBSET_MEMBER: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6088
   PSUBSET x xa = (SUBSET x xa & (EX y::'A::type. IN y xa & ~ IN y x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6089
  by (import hollight PSUBSET_MEMBER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6090
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6091
lemma DELETE_INSERT: "ALL (x::'A::type) (y::'A::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6092
   DELETE (INSERT x s) y = COND (x = y) (DELETE s y) (INSERT x (DELETE s y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6093
  by (import hollight DELETE_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6094
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6095
lemma INSERT_DELETE: "ALL (x::'A::type) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6096
   IN x xa --> INSERT x (DELETE xa x) = xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6097
  by (import hollight INSERT_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6098
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6099
lemma DELETE_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6100
   hollight.INTER (DELETE x xb) xa = DELETE (hollight.INTER x xa) xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6101
  by (import hollight DELETE_INTER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6102
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6103
lemma DISJOINT_DELETE_SYM: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6104
   DISJOINT (DELETE x xb) xa = DISJOINT (DELETE xa xb) x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6105
  by (import hollight DISJOINT_DELETE_SYM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6106
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6107
lemma UNIONS_0: "(op =::('q_40008::type => bool) => ('q_40008::type => bool) => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6108
 ((UNIONS::(('q_40008::type => bool) => bool) => 'q_40008::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6109
   (EMPTY::('q_40008::type => bool) => bool))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6110
 (EMPTY::'q_40008::type => bool)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6111
  by (import hollight UNIONS_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6112
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6113
lemma UNIONS_1: "UNIONS (INSERT (s::'q_40014::type => bool) EMPTY) = s"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6114
  by (import hollight UNIONS_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6115
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6116
lemma UNIONS_2: "UNIONS
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6117
 (INSERT (s::'q_40034::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6118
   (INSERT (t::'q_40034::type => bool) EMPTY)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6119
hollight.UNION s t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6120
  by (import hollight UNIONS_2)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6121
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6122
lemma UNIONS_INSERT: "UNIONS
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6123
 (INSERT (s::'q_40048::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6124
   (u::('q_40048::type => bool) => bool)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6125
hollight.UNION s (UNIONS u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6126
  by (import hollight UNIONS_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6127
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6128
lemma FORALL_IN_UNIONS: "ALL (x::'q_40090::type => bool) xa::('q_40090::type => bool) => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6129
   (ALL xb::'q_40090::type. IN xb (UNIONS xa) --> x xb) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6130
   (ALL (t::'q_40090::type => bool) xb::'q_40090::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6131
       IN t xa & IN xb t --> x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6132
  by (import hollight FORALL_IN_UNIONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6133
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6134
lemma EMPTY_UNIONS: "ALL x::('q_40116::type => bool) => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6135
   (UNIONS x = EMPTY) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6136
   (ALL xa::'q_40116::type => bool. IN xa x --> xa = EMPTY)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6137
  by (import hollight EMPTY_UNIONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6138
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6139
lemma INTERS_0: "(op =::('q_40124::type => bool) => ('q_40124::type => bool) => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6140
 ((INTERS::(('q_40124::type => bool) => bool) => 'q_40124::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6141
   (EMPTY::('q_40124::type => bool) => bool))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6142
 (hollight.UNIV::'q_40124::type => bool)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6143
  by (import hollight INTERS_0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6144
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6145
lemma INTERS_1: "INTERS (INSERT (s::'q_40130::type => bool) EMPTY) = s"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6146
  by (import hollight INTERS_1)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6147
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6148
lemma INTERS_2: "INTERS
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6149
 (INSERT (s::'q_40150::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6150
   (INSERT (t::'q_40150::type => bool) EMPTY)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6151
hollight.INTER s t"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6152
  by (import hollight INTERS_2)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6153
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6154
lemma INTERS_INSERT: "INTERS
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6155
 (INSERT (s::'q_40164::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6156
   (u::('q_40164::type => bool) => bool)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6157
hollight.INTER s (INTERS u)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6158
  by (import hollight INTERS_INSERT)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6159
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6160
lemma IMAGE_CLAUSES: "IMAGE (f::'q_40190::type => 'q_40194::type) EMPTY = EMPTY &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6161
IMAGE f (INSERT (x::'q_40190::type) (s::'q_40190::type => bool)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6162
INSERT (f x) (IMAGE f s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6163
  by (import hollight IMAGE_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6164
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6165
lemma IMAGE_UNION: "ALL (x::'q_40217::type => 'q_40228::type) (xa::'q_40217::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6166
   xb::'q_40217::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6167
   IMAGE x (hollight.UNION xa xb) = hollight.UNION (IMAGE x xa) (IMAGE x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6168
  by (import hollight IMAGE_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6169
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6170
lemma IMAGE_o: "ALL (x::'q_40261::type => 'q_40257::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6171
   (xa::'q_40252::type => 'q_40261::type) xb::'q_40252::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6172
   IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6173
  by (import hollight IMAGE_o)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6174
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6175
lemma IMAGE_SUBSET: "ALL (x::'q_40279::type => 'q_40290::type) (xa::'q_40279::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6176
   xb::'q_40279::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6177
   SUBSET xa xb --> SUBSET (IMAGE x xa) (IMAGE x xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6178
  by (import hollight IMAGE_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6179
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6180
lemma IMAGE_DIFF_INJ: "(ALL (x::'q_40321::type) y::'q_40321::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6181
    (f::'q_40321::type => 'q_40332::type) x = f y --> x = y) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6182
IMAGE f (DIFF (s::'q_40321::type => bool) (t::'q_40321::type => bool)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6183
DIFF (IMAGE f s) (IMAGE f t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6184
  by (import hollight IMAGE_DIFF_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6185
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6186
lemma IMAGE_DELETE_INJ: "(ALL x::'q_40367::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6187
    (f::'q_40367::type => 'q_40366::type) x = f (a::'q_40367::type) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6188
    x = a) -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6189
IMAGE f (DELETE (s::'q_40367::type => bool) a) = DELETE (IMAGE f s) (f a)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6190
  by (import hollight IMAGE_DELETE_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6191
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6192
lemma IMAGE_EQ_EMPTY: "ALL (x::'q_40390::type => 'q_40386::type) xa::'q_40390::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6193
   (IMAGE x xa = EMPTY) = (xa = EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6194
  by (import hollight IMAGE_EQ_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6195
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6196
lemma FORALL_IN_IMAGE: "ALL (x::'q_40426::type => 'q_40425::type) xa::'q_40426::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6197
   (ALL xb::'q_40425::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6198
       IN xb (IMAGE x xa) --> (P::'q_40425::type => bool) xb) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6199
   (ALL xb::'q_40426::type. IN xb xa --> P (x xb))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6200
  by (import hollight FORALL_IN_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6201
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6202
lemma EXISTS_IN_IMAGE: "ALL (x::'q_40462::type => 'q_40461::type) xa::'q_40462::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6203
   (EX xb::'q_40461::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6204
       IN xb (IMAGE x xa) & (P::'q_40461::type => bool) xb) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6205
   (EX xb::'q_40462::type. IN xb xa & P (x xb))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6206
  by (import hollight EXISTS_IN_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6207
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6208
lemma SUBSET_IMAGE: "ALL (f::'A::type => 'B::type) (s::'B::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6209
   SUBSET s (IMAGE f t) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6210
   (EX x::'A::type => bool. SUBSET x t & s = IMAGE f x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6211
  by (import hollight SUBSET_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6212
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6213
lemma IMAGE_CONST: "ALL (s::'q_40548::type => bool) c::'q_40553::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6214
   IMAGE (%x::'q_40548::type. c) s = COND (s = EMPTY) EMPTY (INSERT c EMPTY)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6215
  by (import hollight IMAGE_CONST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6216
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6217
lemma SIMPLE_IMAGE: "ALL (x::'q_40581::type => 'q_40585::type) xa::'q_40581::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6218
   GSPEC
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6219
    (%u::'q_40585::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6220
        EX xb::'q_40581::type. SETSPEC u (IN xb xa) (x xb)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6221
   IMAGE x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6222
  by (import hollight SIMPLE_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6223
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6224
lemma EMPTY_GSPEC: "GSPEC (%u::'q_40602::type. Ex (SETSPEC u False)) = EMPTY"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6225
  by (import hollight EMPTY_GSPEC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6226
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6227
lemma IN_ELIM_PAIR_THM: "ALL (x::'q_40648::type => 'q_40647::type => bool) (xa::'q_40648::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6228
   xb::'q_40647::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6229
   IN (xa, xb)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6230
    (GSPEC
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6231
      (%xa::'q_40648::type * 'q_40647::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6232
          EX (xb::'q_40648::type) y::'q_40647::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6233
             SETSPEC xa (x xb y) (xb, y))) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6234
   x xa xb"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6235
  by (import hollight IN_ELIM_PAIR_THM)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6236
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6237
lemma FINITE_INDUCT_STRONG: "ALL P::('A::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6238
   P EMPTY &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6239
   (ALL (x::'A::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6240
       P s & ~ IN x s & FINITE s --> P (INSERT x s)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6241
   (ALL s::'A::type => bool. FINITE s --> P s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6242
  by (import hollight FINITE_INDUCT_STRONG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6243
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6244
lemma FINITE_SUBSET: "ALL (x::'A::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6245
   FINITE t & SUBSET x t --> FINITE x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6246
  by (import hollight FINITE_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6247
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6248
lemma FINITE_UNION_IMP: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6249
   FINITE x & FINITE xa --> FINITE (hollight.UNION x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6250
  by (import hollight FINITE_UNION_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6251
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6252
lemma FINITE_UNION: "ALL (s::'A::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6253
   FINITE (hollight.UNION s t) = (FINITE s & FINITE t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6254
  by (import hollight FINITE_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6255
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6256
lemma FINITE_INTER: "ALL (s::'A::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6257
   FINITE s | FINITE t --> FINITE (hollight.INTER s t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6258
  by (import hollight FINITE_INTER)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6259
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6260
lemma FINITE_INSERT: "ALL (s::'A::type => bool) x::'A::type. FINITE (INSERT x s) = FINITE s"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6261
  by (import hollight FINITE_INSERT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6262
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6263
lemma FINITE_DELETE_IMP: "ALL (s::'A::type => bool) x::'A::type. FINITE s --> FINITE (DELETE s x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6264
  by (import hollight FINITE_DELETE_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6265
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6266
lemma FINITE_DELETE: "ALL (s::'A::type => bool) x::'A::type. FINITE (DELETE s x) = FINITE s"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6267
  by (import hollight FINITE_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6268
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6269
lemma FINITE_UNIONS: "ALL s::('q_40983::type => bool) => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6270
   FINITE s -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6271
   FINITE (UNIONS s) = (ALL t::'q_40983::type => bool. IN t s --> FINITE t)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6272
  by (import hollight FINITE_UNIONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6273
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6274
lemma FINITE_IMAGE_EXPAND: "ALL (f::'A::type => 'B::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6275
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6276
   FINITE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6277
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6278
      (%u::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6279
          EX y::'B::type. SETSPEC u (EX x::'A::type. IN x s & y = f x) y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6280
  by (import hollight FINITE_IMAGE_EXPAND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6281
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6282
lemma FINITE_IMAGE: "ALL (x::'A::type => 'B::type) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6283
   FINITE xa --> FINITE (IMAGE x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6284
  by (import hollight FINITE_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6285
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6286
lemma FINITE_IMAGE_INJ_GENERAL: "ALL (f::'A::type => 'B::type) (x::'B::type => bool) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6287
   (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:
diff changeset
  6288
   FINITE x -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6289
   FINITE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6290
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6291
      (%u::'A::type. EX xa::'A::type. SETSPEC u (IN xa s & IN (f xa) x) xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6292
  by (import hollight FINITE_IMAGE_INJ_GENERAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6293
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6294
lemma FINITE_IMAGE_INJ: "ALL (f::'A::type => 'B::type) A::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6295
   (ALL (x::'A::type) y::'A::type. f x = f y --> x = y) & FINITE A -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6296
   FINITE (GSPEC (%u::'A::type. EX x::'A::type. SETSPEC u (IN (f x) A) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6297
  by (import hollight FINITE_IMAGE_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6298
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6299
lemma INFINITE_IMAGE_INJ: "ALL f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6300
   (ALL (x::'A::type) y::'A::type. f x = f y --> x = y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6301
   (ALL s::'A::type => bool. INFINITE s --> INFINITE (IMAGE f s))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6302
  by (import hollight INFINITE_IMAGE_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6303
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6304
lemma INFINITE_NONEMPTY: "ALL s::'q_41466::type => bool. INFINITE s --> s ~= EMPTY"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6305
  by (import hollight INFINITE_NONEMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6306
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6307
lemma INFINITE_DIFF_FINITE: "ALL (s::'A::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6308
   INFINITE s & FINITE t --> INFINITE (DIFF s t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6309
  by (import hollight INFINITE_DIFF_FINITE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6310
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6311
lemma FINITE_SUBSET_IMAGE: "ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6312
   (FINITE t & SUBSET t (IMAGE f s)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6313
   (EX x::'A::type => bool. FINITE x & SUBSET x s & t = IMAGE f x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6314
  by (import hollight FINITE_SUBSET_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6315
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6316
lemma FINITE_SUBSET_IMAGE_IMP: "ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6317
   FINITE t & SUBSET t (IMAGE f s) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6318
   (EX s'::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6319
       FINITE s' & SUBSET s' s & SUBSET t (IMAGE f s'))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6320
  by (import hollight FINITE_SUBSET_IMAGE_IMP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6321
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6322
lemma FINITE_SUBSETS: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6323
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6324
   FINITE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6325
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6326
      (%u::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6327
          EX t::'A::type => bool. SETSPEC u (SUBSET t s) t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6328
  by (import hollight FINITE_SUBSETS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6329
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6330
lemma FINITE_DIFF: "ALL (s::'q_41764::type => bool) t::'q_41764::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6331
   FINITE s --> FINITE (DIFF s t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6332
  by (import hollight FINITE_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6333
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6334
definition FINREC :: "('q_41824 => 'q_41823 => 'q_41823)
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6335
=> 'q_41823 => ('q_41824 => bool) => 'q_41823 => nat => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6336
  "FINREC ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6337
SOME FINREC::('q_41824::type => 'q_41823::type => 'q_41823::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6338
             => 'q_41823::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6339
                => ('q_41824::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6340
                   => 'q_41823::type => nat => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6341
   (ALL (f::'q_41824::type => 'q_41823::type => 'q_41823::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6342
       (s::'q_41824::type => bool) (a::'q_41823::type) b::'q_41823::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6343
       FINREC f b s a 0 = (s = EMPTY & a = b)) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6344
   (ALL (b::'q_41823::type) (s::'q_41824::type => bool) (n::nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6345
       (a::'q_41823::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6346
       f::'q_41824::type => 'q_41823::type => 'q_41823::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6347
       FINREC f b s a (Suc n) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6348
       (EX (x::'q_41824::type) c::'q_41823::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6349
           IN x s & FINREC f b (DELETE s x) c n & a = f x c))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6350
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6351
lemma DEF_FINREC: "FINREC =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6352
(SOME FINREC::('q_41824::type => 'q_41823::type => 'q_41823::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6353
              => 'q_41823::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6354
                 => ('q_41824::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6355
                    => 'q_41823::type => nat => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6356
    (ALL (f::'q_41824::type => 'q_41823::type => 'q_41823::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6357
        (s::'q_41824::type => bool) (a::'q_41823::type) b::'q_41823::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6358
        FINREC f b s a 0 = (s = EMPTY & a = b)) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6359
    (ALL (b::'q_41823::type) (s::'q_41824::type => bool) (n::nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6360
        (a::'q_41823::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6361
        f::'q_41824::type => 'q_41823::type => 'q_41823::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6362
        FINREC f b s a (Suc n) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6363
        (EX (x::'q_41824::type) c::'q_41823::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6364
            IN x s & FINREC f b (DELETE s x) c n & a = f x c)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6365
  by (import hollight DEF_FINREC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6366
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6367
lemma FINREC_1_LEMMA: "ALL (x::'q_41869::type => 'q_41868::type => 'q_41868::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6368
   (xa::'q_41868::type) (xb::'q_41869::type => bool) xc::'q_41868::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6369
   FINREC x xa xb xc (Suc 0) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6370
   (EX xd::'q_41869::type. xb = INSERT xd EMPTY & xc = x xd xa)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6371
  by (import hollight FINREC_1_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6372
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6373
lemma FINREC_SUC_LEMMA: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6374
   (ALL (x::'A::type) (y::'A::type) s::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6375
       x ~= y --> f x (f y s) = f y (f x s)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6376
   (ALL (n::nat) (s::'A::type => bool) z::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6377
       FINREC f b s z (Suc n) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6378
       (ALL x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6379
           IN x s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6380
           (EX w::'B::type. FINREC f b (DELETE s x) w n & z = f x w)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6381
  by (import hollight FINREC_SUC_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6382
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6383
lemma FINREC_UNIQUE_LEMMA: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6384
   (ALL (x::'A::type) (y::'A::type) s::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6385
       x ~= y --> f x (f y s) = f y (f x s)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6386
   (ALL (n1::nat) (n2::nat) (s::'A::type => bool) (a1::'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6387
       a2::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6388
       FINREC f b s a1 n1 & FINREC f b s a2 n2 --> a1 = a2 & n1 = n2)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6389
  by (import hollight FINREC_UNIQUE_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6390
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6391
lemma FINREC_EXISTS_LEMMA: "ALL (f::'A::type => 'B::type => 'B::type) (b::'B::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6392
   FINITE s --> (EX a::'B::type. Ex (FINREC f b s a))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6393
  by (import hollight FINREC_EXISTS_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6394
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6395
lemma FINREC_FUN_LEMMA: "ALL (P::'A::type => bool) R::'A::type => 'B::type => 'C::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6396
   (ALL s::'A::type. P s --> (EX a::'B::type. Ex (R s a))) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6397
   (ALL (n1::'C::type) (n2::'C::type) (s::'A::type) (a1::'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6398
       a2::'B::type. R s a1 n1 & R s a2 n2 --> a1 = a2 & n1 = n2) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6399
   (EX x::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6400
       ALL (s::'A::type) a::'B::type. P s --> Ex (R s a) = (x s = a))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6401
  by (import hollight FINREC_FUN_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6402
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6403
lemma FINREC_FUN: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6404
   (ALL (x::'A::type) (y::'A::type) s::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6405
       x ~= y --> f x (f y s) = f y (f x s)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6406
   (EX g::('A::type => bool) => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6407
       g EMPTY = b &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6408
       (ALL (s::'A::type => bool) x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6409
           FINITE s & IN x s --> g s = f x (g (DELETE s x))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6410
  by (import hollight FINREC_FUN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6411
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6412
lemma SET_RECURSION_LEMMA: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6413
   (ALL (x::'A::type) (y::'A::type) s::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6414
       x ~= y --> f x (f y s) = f y (f x s)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6415
   (EX g::('A::type => bool) => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6416
       g EMPTY = b &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6417
       (ALL (x::'A::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6418
           FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (f x (g s))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6419
  by (import hollight SET_RECURSION_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6420
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6421
definition ITSET :: "('q_42525 => 'q_42524 => 'q_42524)
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6422
=> ('q_42525 => bool) => 'q_42524 => 'q_42524" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6423
  "ITSET ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6424
%(u::'q_42525::type => 'q_42524::type => 'q_42524::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6425
   (ua::'q_42525::type => bool) ub::'q_42524::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6426
   (SOME g::('q_42525::type => bool) => 'q_42524::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6427
       g EMPTY = ub &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6428
       (ALL (x::'q_42525::type) s::'q_42525::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6429
           FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (u x (g s))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6430
    ua"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6431
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6432
lemma DEF_ITSET: "ITSET =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6433
(%(u::'q_42525::type => 'q_42524::type => 'q_42524::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6434
    (ua::'q_42525::type => bool) ub::'q_42524::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6435
    (SOME g::('q_42525::type => bool) => 'q_42524::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6436
        g EMPTY = ub &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6437
        (ALL (x::'q_42525::type) s::'q_42525::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6438
            FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (u x (g s))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6439
     ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6440
  by (import hollight DEF_ITSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6441
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6442
lemma FINITE_RECURSION: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6443
   (ALL (x::'A::type) (y::'A::type) s::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6444
       x ~= y --> f x (f y s) = f y (f x s)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6445
   ITSET f EMPTY b = b &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6446
   (ALL (x::'A::type) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6447
       FINITE xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6448
       ITSET f (INSERT x xa) b =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6449
       COND (IN x xa) (ITSET f xa b) (f x (ITSET f xa b)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6450
  by (import hollight FINITE_RECURSION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6451
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6452
lemma FINITE_RECURSION_DELETE: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6453
   (ALL (x::'A::type) (y::'A::type) s::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6454
       x ~= y --> f x (f y s) = f y (f x s)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6455
   ITSET f EMPTY b = b &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6456
   (ALL (x::'A::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6457
       FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6458
       ITSET f s b =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6459
       COND (IN x s) (f x (ITSET f (DELETE s x) b))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6460
        (ITSET f (DELETE s x) b))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6461
  by (import hollight FINITE_RECURSION_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6462
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6463
lemma ITSET_EQ: "ALL (x::'q_42830::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6464
   (xa::'q_42830::type => 'q_42831::type => 'q_42831::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6465
   (xb::'q_42830::type => 'q_42831::type => 'q_42831::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6466
   xc::'q_42831::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6467
   FINITE x &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6468
   (ALL xc::'q_42830::type. IN xc x --> xa xc = xb xc) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6469
   (ALL (x::'q_42830::type) (y::'q_42830::type) s::'q_42831::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6470
       x ~= y --> xa x (xa y s) = xa y (xa x s)) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6471
   (ALL (x::'q_42830::type) (y::'q_42830::type) s::'q_42831::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6472
       x ~= y --> xb x (xb y s) = xb y (xb x s)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6473
   ITSET xa x xc = ITSET xb x xc"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6474
  by (import hollight ITSET_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6475
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6476
lemma SUBSET_RESTRICT: "ALL (x::'q_42864::type => bool) xa::'q_42864::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6477
   SUBSET
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6478
    (GSPEC
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6479
      (%u::'q_42864::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6480
          EX xb::'q_42864::type. SETSPEC u (IN xb x & xa xb) xb))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6481
    x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6482
  by (import hollight SUBSET_RESTRICT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6483
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6484
lemma FINITE_RESTRICT: "ALL (s::'A::type => bool) p::'q_42882::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6485
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6486
   FINITE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6487
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6488
      (%u::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6489
          EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6490
  by (import hollight FINITE_RESTRICT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6491
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6492
definition CARD :: "('q_42918 => bool) => nat" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6493
  "CARD == %u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6494
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6495
lemma DEF_CARD: "CARD = (%u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6496
  by (import hollight DEF_CARD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6497
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6498
lemma CARD_CLAUSES: "(op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6499
 ((op =::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6500
   ((CARD::('A::type => bool) => nat) (EMPTY::'A::type => bool)) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6501
 ((All::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6502
   (%x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6503
       (All::(('A::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6504
        (%s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6505
            (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6506
             ((FINITE::('A::type => bool) => bool) s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6507
             ((op =::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6508
               ((CARD::('A::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6509
                 ((INSERT::'A::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6510
                           => ('A::type => bool) => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6511
                   x s))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6512
               ((COND::bool => nat => nat => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6513
                 ((IN::'A::type => ('A::type => bool) => bool) x s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6514
                 ((CARD::('A::type => bool) => nat) s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6515
                 ((Suc::nat => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6516
                   ((CARD::('A::type => bool) => nat) s)))))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6517
  by (import hollight CARD_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6518
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6519
lemma CARD_UNION: "ALL (x::'A::type => bool) xa::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6520
   FINITE x & FINITE xa & hollight.INTER x xa = EMPTY -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6521
   CARD (hollight.UNION x xa) = CARD x + CARD xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6522
  by (import hollight CARD_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6523
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6524
lemma CARD_DELETE: "ALL (x::'A::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6525
   FINITE s -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6526
   CARD (DELETE s x) = COND (IN x s) (CARD s - NUMERAL_BIT1 0) (CARD s)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6527
  by (import hollight CARD_DELETE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6528
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6529
lemma CARD_UNION_EQ: "ALL (s::'q_43163::type => bool) (t::'q_43163::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6530
   u::'q_43163::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6531
   FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6532
   CARD s + CARD t = CARD u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6533
  by (import hollight CARD_UNION_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6534
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6535
definition HAS_SIZE :: "('q_43199 => bool) => nat => bool" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6536
  "HAS_SIZE == %(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6537
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6538
lemma DEF_HAS_SIZE: "HAS_SIZE = (%(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6539
  by (import hollight DEF_HAS_SIZE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6540
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6541
lemma HAS_SIZE_CARD: "ALL (x::'q_43218::type => bool) xa::nat. HAS_SIZE x xa --> CARD x = xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6542
  by (import hollight HAS_SIZE_CARD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6543
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6544
lemma HAS_SIZE_0: "ALL (s::'A::type => bool) n::'q_43234::type. HAS_SIZE s 0 = (s = EMPTY)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6545
  by (import hollight HAS_SIZE_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6546
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6547
lemma HAS_SIZE_SUC: "ALL (s::'A::type => bool) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6548
   HAS_SIZE s (Suc n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6549
   (s ~= EMPTY & (ALL x::'A::type. IN x s --> HAS_SIZE (DELETE s x) n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6550
  by (import hollight HAS_SIZE_SUC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6551
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6552
lemma HAS_SIZE_UNION: "ALL (x::'q_43356::type => bool) (xa::'q_43356::type => bool) (xb::nat)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6553
   xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6554
   HAS_SIZE x xb & HAS_SIZE xa xc & DISJOINT x xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6555
   HAS_SIZE (hollight.UNION x xa) (xb + xc)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6556
  by (import hollight HAS_SIZE_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6557
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6558
lemma HAS_SIZE_UNIONS: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) (xb::nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6559
   xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6560
   HAS_SIZE x xb &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6561
   (ALL xb::'A::type. IN xb x --> HAS_SIZE (xa xb) xc) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6562
   (ALL (xb::'A::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6563
       IN xb x & IN y x & xb ~= y --> DISJOINT (xa xb) (xa y)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6564
   HAS_SIZE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6565
    (UNIONS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6566
      (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6567
        (%u::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6568
            EX xb::'A::type. SETSPEC u (IN xb x) (xa xb))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6569
    (xb * xc)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6570
  by (import hollight HAS_SIZE_UNIONS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6571
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6572
lemma HAS_SIZE_CLAUSES: "HAS_SIZE (s::'q_43604::type => bool) 0 = (s = EMPTY) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6573
HAS_SIZE s (Suc (n::nat)) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6574
(EX (a::'q_43604::type) t::'q_43604::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6575
    HAS_SIZE t n & ~ IN a t & s = INSERT a t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6576
  by (import hollight HAS_SIZE_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6577
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6578
lemma CARD_SUBSET_EQ: "ALL (a::'A::type => bool) b::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6579
   FINITE b & SUBSET a b & CARD a = CARD b --> a = b"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6580
  by (import hollight CARD_SUBSET_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6581
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6582
lemma CARD_SUBSET: "ALL (a::'A::type => bool) b::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6583
   SUBSET a b & FINITE b --> <= (CARD a) (CARD b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6584
  by (import hollight CARD_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6585
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6586
lemma CARD_SUBSET_LE: "ALL (a::'A::type => bool) b::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6587
   FINITE b & SUBSET a b & <= (CARD b) (CARD a) --> a = b"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6588
  by (import hollight CARD_SUBSET_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6589
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6590
lemma CARD_EQ_0: "ALL s::'q_43920::type => bool. FINITE s --> (CARD s = 0) = (s = EMPTY)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6591
  by (import hollight CARD_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6592
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6593
lemma CARD_PSUBSET: "ALL (a::'A::type => bool) b::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6594
   PSUBSET a b & FINITE b --> < (CARD a) (CARD b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6595
  by (import hollight CARD_PSUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6596
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6597
lemma CARD_UNION_LE: "ALL (s::'A::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6598
   FINITE s & FINITE t --> <= (CARD (hollight.UNION s t)) (CARD s + CARD t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6599
  by (import hollight CARD_UNION_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6600
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6601
lemma CARD_UNIONS_LE: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) (xb::nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6602
   xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6603
   HAS_SIZE x xb &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6604
   (ALL xb::'A::type. IN xb x --> FINITE (xa xb) & <= (CARD (xa xb)) xc) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6605
   <= (CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6606
        (UNIONS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6607
          (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6608
            (%u::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6609
                EX xb::'A::type. SETSPEC u (IN xb x) (xa xb)))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6610
    (xb * xc)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6611
  by (import hollight CARD_UNIONS_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6612
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6613
lemma CARD_IMAGE_INJ: "ALL (f::'A::type => 'B::type) x::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6614
   (ALL (xa::'A::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6615
       IN xa x & IN y x & f xa = f y --> xa = y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6616
   FINITE x -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6617
   CARD (IMAGE f x) = CARD x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6618
  by (import hollight CARD_IMAGE_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6619
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6620
lemma HAS_SIZE_IMAGE_INJ: "ALL (x::'A::type => 'B::type) (xa::'A::type => bool) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6621
   (ALL (xb::'A::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6622
       IN xb xa & IN y xa & x xb = x y --> xb = y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6623
   HAS_SIZE xa xb -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6624
   HAS_SIZE (IMAGE x xa) xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6625
  by (import hollight HAS_SIZE_IMAGE_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6626
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6627
lemma CARD_IMAGE_LE: "ALL (f::'A::type => 'B::type) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6628
   FINITE s --> <= (CARD (IMAGE f s)) (CARD s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6629
  by (import hollight CARD_IMAGE_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6630
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6631
lemma CHOOSE_SUBSET: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6632
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6633
   (ALL n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6634
       <= n (CARD s) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6635
       (EX t::'A::type => bool. SUBSET t s & HAS_SIZE t n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6636
  by (import hollight CHOOSE_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6637
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6638
lemma HAS_SIZE_PRODUCT_DEPENDENT: "ALL (x::'A::type => bool) (xa::nat) (xb::'A::type => 'B::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6639
   xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6640
   HAS_SIZE x xa & (ALL xa::'A::type. IN xa x --> HAS_SIZE (xb xa) xc) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6641
   HAS_SIZE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6642
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6643
      (%u::'A::type * 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6644
          EX (xa::'A::type) y::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6645
             SETSPEC u (IN xa x & IN y (xb xa)) (xa, y)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6646
    (xa * xc)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6647
  by (import hollight HAS_SIZE_PRODUCT_DEPENDENT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6648
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6649
lemma FINITE_PRODUCT_DEPENDENT: "ALL (x::'A::type => bool) xa::'A::type => 'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6650
   FINITE x & (ALL xb::'A::type. IN xb x --> FINITE (xa xb)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6651
   FINITE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6652
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6653
      (%u::'A::type * 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6654
          EX (xb::'A::type) y::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6655
             SETSPEC u (IN xb x & IN y (xa xb)) (xb, y)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6656
  by (import hollight FINITE_PRODUCT_DEPENDENT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6657
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6658
lemma FINITE_PRODUCT: "ALL (x::'A::type => bool) xa::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6659
   FINITE x & FINITE xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6660
   FINITE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6661
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6662
      (%u::'A::type * 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6663
          EX (xb::'A::type) y::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6664
             SETSPEC u (IN xb x & IN y xa) (xb, y)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6665
  by (import hollight FINITE_PRODUCT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6666
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6667
lemma CARD_PRODUCT: "ALL (s::'A::type => bool) t::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6668
   FINITE s & FINITE t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6669
   CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6670
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6671
      (%u::'A::type * 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6672
          EX (x::'A::type) y::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6673
             SETSPEC u (IN x s & IN y t) (x, y))) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6674
   CARD s * CARD t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6675
  by (import hollight CARD_PRODUCT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6676
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6677
lemma HAS_SIZE_PRODUCT: "ALL (x::'A::type => bool) (xa::nat) (xb::'B::type => bool) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6678
   HAS_SIZE x xa & HAS_SIZE xb xc -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6679
   HAS_SIZE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6680
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6681
      (%u::'A::type * 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6682
          EX (xa::'A::type) y::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6683
             SETSPEC u (IN xa x & IN y xb) (xa, y)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6684
    (xa * xc)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6685
  by (import hollight HAS_SIZE_PRODUCT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6686
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6687
lemma HAS_SIZE_FUNSPACE: "ALL (d::'B::type) (n::nat) (t::'B::type => bool) (m::nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6688
   s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6689
   HAS_SIZE s m & HAS_SIZE t n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6690
   HAS_SIZE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6691
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6692
      (%u::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6693
          EX f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6694
             SETSPEC u
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6695
              ((ALL x::'A::type. IN x s --> IN (f x) t) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6696
               (ALL x::'A::type. ~ IN x s --> f x = d))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6697
              f))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6698
    (EXP n m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6699
  by (import hollight HAS_SIZE_FUNSPACE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6700
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6701
lemma CARD_FUNSPACE: "ALL (s::'q_45275::type => bool) t::'q_45272::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6702
   FINITE s & FINITE t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6703
   CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6704
    (GSPEC
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6705
      (%u::'q_45275::type => 'q_45272::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6706
          EX f::'q_45275::type => 'q_45272::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6707
             SETSPEC u
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6708
              ((ALL x::'q_45275::type. IN x s --> IN (f x) t) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6709
               (ALL x::'q_45275::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6710
                   ~ IN x s --> f x = (d::'q_45272::type)))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6711
              f)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6712
   EXP (CARD t) (CARD s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6713
  by (import hollight CARD_FUNSPACE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6714
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6715
lemma FINITE_FUNSPACE: "ALL (s::'q_45341::type => bool) t::'q_45338::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6716
   FINITE s & FINITE t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6717
   FINITE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6718
    (GSPEC
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6719
      (%u::'q_45341::type => 'q_45338::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6720
          EX f::'q_45341::type => 'q_45338::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6721
             SETSPEC u
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6722
              ((ALL x::'q_45341::type. IN x s --> IN (f x) t) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6723
               (ALL x::'q_45341::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6724
                   ~ IN x s --> f x = (d::'q_45338::type)))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6725
              f))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6726
  by (import hollight FINITE_FUNSPACE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6727
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6728
lemma HAS_SIZE_POWERSET: "ALL (s::'A::type => bool) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6729
   HAS_SIZE s n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6730
   HAS_SIZE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6731
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6732
      (%u::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6733
          EX t::'A::type => bool. SETSPEC u (SUBSET t s) t))
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6734
    (EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6735
  by (import hollight HAS_SIZE_POWERSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6736
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6737
lemma CARD_POWERSET: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6738
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6739
   CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6740
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6741
      (%u::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6742
          EX t::'A::type => bool. SETSPEC u (SUBSET t s) t)) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6743
   EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) (CARD s)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6744
  by (import hollight CARD_POWERSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6745
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6746
lemma FINITE_POWERSET: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6747
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6748
   FINITE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6749
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6750
      (%u::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6751
          EX t::'A::type => bool. SETSPEC u (SUBSET t s) t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6752
  by (import hollight FINITE_POWERSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6753
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6754
lemma CARD_GE_REFL: "ALL s::'A::type => bool. CARD_GE s s"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6755
  by (import hollight CARD_GE_REFL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6756
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6757
lemma CARD_GE_TRANS: "ALL (s::'A::type => bool) (t::'B::type => bool) u::'C::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6758
   CARD_GE s t & CARD_GE t u --> CARD_GE s u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6759
  by (import hollight CARD_GE_TRANS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6760
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6761
lemma FINITE_HAS_SIZE_LEMMA: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6762
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6763
   (EX n::nat. CARD_GE (GSPEC (%u::nat. EX x::nat. SETSPEC u (< x n) x)) s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6764
  by (import hollight FINITE_HAS_SIZE_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6765
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6766
lemma HAS_SIZE_NUMSEG_LT: "ALL n::nat. HAS_SIZE (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m n) m)) n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6767
  by (import hollight HAS_SIZE_NUMSEG_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6768
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6769
lemma CARD_NUMSEG_LT: "ALL x::nat. CARD (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m x) m)) = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6770
  by (import hollight CARD_NUMSEG_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6771
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6772
lemma FINITE_NUMSEG_LT: "ALL x::nat. FINITE (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m x) m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6773
  by (import hollight FINITE_NUMSEG_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6774
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6775
lemma HAS_SIZE_NUMSEG_LE: "ALL x::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6776
   HAS_SIZE (GSPEC (%xa::nat. EX xb::nat. SETSPEC xa (<= xb x) xb))
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6777
    (x + NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6778
  by (import hollight HAS_SIZE_NUMSEG_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6779
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6780
lemma FINITE_NUMSEG_LE: "ALL x::nat. FINITE (GSPEC (%u::nat. EX m::nat. SETSPEC u (<= m x) m))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6781
  by (import hollight FINITE_NUMSEG_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6782
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6783
lemma CARD_NUMSEG_LE: "ALL x::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6784
   CARD (GSPEC (%u::nat. EX m::nat. SETSPEC u (<= m x) m)) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6785
   x + NUMERAL_BIT1 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6786
  by (import hollight CARD_NUMSEG_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6787
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6788
lemma num_FINITE: "ALL s::nat => bool. FINITE s = (EX a::nat. ALL x::nat. IN x s --> <= x a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6789
  by (import hollight num_FINITE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6790
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6791
lemma num_FINITE_AVOID: "ALL s::nat => bool. FINITE s --> (EX a::nat. ~ IN a s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6792
  by (import hollight num_FINITE_AVOID)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6793
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6794
lemma num_INFINITE: "(INFINITE::(nat => bool) => bool) (hollight.UNIV::nat => bool)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6795
  by (import hollight num_INFINITE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6796
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6797
lemma HAS_SIZE_INDEX: "ALL (x::'A::type => bool) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6798
   HAS_SIZE x n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6799
   (EX f::nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6800
       (ALL m::nat. < m n --> IN (f m) x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6801
       (ALL xa::'A::type. IN xa x --> (EX! m::nat. < m n & f m = xa)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6802
  by (import hollight HAS_SIZE_INDEX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6803
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6804
definition set_of_list :: "'q_45968 hollight.list => 'q_45968 => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6805
  "set_of_list ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6806
SOME set_of_list::'q_45968::type hollight.list => 'q_45968::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6807
   set_of_list NIL = EMPTY &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6808
   (ALL (h::'q_45968::type) t::'q_45968::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6809
       set_of_list (CONS h t) = INSERT h (set_of_list t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6810
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6811
lemma DEF_set_of_list: "set_of_list =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6812
(SOME set_of_list::'q_45968::type hollight.list => 'q_45968::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6813
    set_of_list NIL = EMPTY &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6814
    (ALL (h::'q_45968::type) t::'q_45968::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6815
        set_of_list (CONS h t) = INSERT h (set_of_list t)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6816
  by (import hollight DEF_set_of_list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6817
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6818
definition list_of_set :: "('q_45986 => bool) => 'q_45986 hollight.list" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6819
  "list_of_set ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6820
%u::'q_45986::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6821
   SOME l::'q_45986::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6822
      set_of_list l = u & LENGTH l = CARD u"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6823
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6824
lemma DEF_list_of_set: "list_of_set =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6825
(%u::'q_45986::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6826
    SOME l::'q_45986::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6827
       set_of_list l = u & LENGTH l = CARD u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6828
  by (import hollight DEF_list_of_set)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6829
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6830
lemma LIST_OF_SET_PROPERTIES: "ALL x::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6831
   FINITE x -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6832
   set_of_list (list_of_set x) = x & LENGTH (list_of_set x) = CARD x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6833
  by (import hollight LIST_OF_SET_PROPERTIES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6834
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6835
lemma SET_OF_LIST_OF_SET: "ALL s::'q_46035::type => bool. FINITE s --> set_of_list (list_of_set s) = s"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6836
  by (import hollight SET_OF_LIST_OF_SET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6837
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6838
lemma LENGTH_LIST_OF_SET: "ALL s::'q_46051::type => bool. FINITE s --> LENGTH (list_of_set s) = CARD s"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6839
  by (import hollight LENGTH_LIST_OF_SET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6840
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6841
lemma MEM_LIST_OF_SET: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6842
   FINITE s --> (ALL x::'A::type. MEM x (list_of_set s) = IN x s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6843
  by (import hollight MEM_LIST_OF_SET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6844
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6845
lemma FINITE_SET_OF_LIST: "ALL l::'q_46096::type hollight.list. FINITE (set_of_list l)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6846
  by (import hollight FINITE_SET_OF_LIST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6847
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6848
lemma IN_SET_OF_LIST: "ALL (x::'q_46114::type) l::'q_46114::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6849
   IN x (set_of_list l) = MEM x l"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6850
  by (import hollight IN_SET_OF_LIST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6851
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6852
lemma SET_OF_LIST_APPEND: "ALL (x::'q_46139::type hollight.list) xa::'q_46139::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6853
   set_of_list (APPEND x xa) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6854
   hollight.UNION (set_of_list x) (set_of_list xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6855
  by (import hollight SET_OF_LIST_APPEND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6856
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6857
definition pairwise :: "('q_46198 => 'q_46198 => bool) => ('q_46198 => bool) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6858
  "pairwise ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6859
%(u::'q_46198::type => 'q_46198::type => bool) ua::'q_46198::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6860
   ALL (x::'q_46198::type) y::'q_46198::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6861
      IN x ua & IN y ua & x ~= y --> u x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6862
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6863
lemma DEF_pairwise: "pairwise =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6864
(%(u::'q_46198::type => 'q_46198::type => bool) ua::'q_46198::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6865
    ALL (x::'q_46198::type) y::'q_46198::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6866
       IN x ua & IN y ua & x ~= y --> u x y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6867
  by (import hollight DEF_pairwise)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6868
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6869
definition PAIRWISE :: "('q_46220 => 'q_46220 => bool) => 'q_46220 hollight.list => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6870
  "PAIRWISE ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6871
SOME PAIRWISE::('q_46220::type => 'q_46220::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6872
               => 'q_46220::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6873
   (ALL r::'q_46220::type => 'q_46220::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6874
       PAIRWISE r NIL = True) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6875
   (ALL (h::'q_46220::type) (r::'q_46220::type => 'q_46220::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6876
       t::'q_46220::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6877
       PAIRWISE r (CONS h t) = (ALL_list (r h) t & PAIRWISE r t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6878
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6879
lemma DEF_PAIRWISE: "PAIRWISE =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6880
(SOME PAIRWISE::('q_46220::type => 'q_46220::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6881
                => 'q_46220::type hollight.list => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6882
    (ALL r::'q_46220::type => 'q_46220::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6883
        PAIRWISE r NIL = True) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6884
    (ALL (h::'q_46220::type) (r::'q_46220::type => 'q_46220::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  6885
        t::'q_46220::type hollight.list.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6886
        PAIRWISE r (CONS h t) = (ALL_list (r h) t & PAIRWISE r t)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6887
  by (import hollight DEF_PAIRWISE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6888
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6889
typedef (open) ('A) finite_image = "(Collect::('A::type => bool) => 'A::type set)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6890
 (%x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6891
     (op |::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6892
      ((op =::'A::type => 'A::type => bool) x
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6893
        ((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6894
      ((FINITE::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6895
        (hollight.UNIV::'A::type => bool)))"  morphisms "dest_finite_image" "mk_finite_image"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6896
  apply (rule light_ex_imp_nonempty[where t="(Eps::('A::type => bool) => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6897
 (%x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6898
     (op |::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6899
      ((op =::'A::type => 'A::type => bool) x
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6900
        ((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6901
      ((FINITE::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6902
        (hollight.UNIV::'A::type => bool)))"])
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6903
  by (import hollight TYDEF_finite_image)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6904
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6905
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6906
  dest_finite_image :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6907
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6908
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6909
  mk_finite_image :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6910
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6911
lemmas "TYDEF_finite_image_@intern" = typedef_hol2hollight 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6912
  [where a="a :: 'A finite_image" and r=r ,
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6913
   OF type_definition_finite_image]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6914
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6915
lemma FINITE_IMAGE_IMAGE: "(op =::('A::type finite_image => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6916
       => ('A::type finite_image => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6917
 (hollight.UNIV::'A::type finite_image => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6918
 ((IMAGE::('A::type => 'A::type finite_image)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6919
          => ('A::type => bool) => 'A::type finite_image => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6920
   (mk_finite_image::'A::type => 'A::type finite_image)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6921
   ((COND::bool
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6922
           => ('A::type => bool) => ('A::type => bool) => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6923
     ((FINITE::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6924
       (hollight.UNIV::'A::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6925
     (hollight.UNIV::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6926
     ((INSERT::'A::type => ('A::type => bool) => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6927
       ((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6928
       (EMPTY::'A::type => bool))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6929
  by (import hollight FINITE_IMAGE_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6930
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6931
definition dimindex :: "('A => bool) => nat" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6932
  "(op ==::(('A::type => bool) => nat) => (('A::type => bool) => nat) => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6933
 (dimindex::('A::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6934
 (%u::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6935
     (COND::bool => nat => nat => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6936
      ((FINITE::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6937
        (hollight.UNIV::'A::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6938
      ((CARD::('A::type => bool) => nat) (hollight.UNIV::'A::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6939
      ((NUMERAL_BIT1::nat => nat) (0::nat)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6940
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6941
lemma DEF_dimindex: "(op =::(('A::type => bool) => nat) => (('A::type => bool) => nat) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6942
 (dimindex::('A::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6943
 (%u::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6944
     (COND::bool => nat => nat => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6945
      ((FINITE::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6946
        (hollight.UNIV::'A::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6947
      ((CARD::('A::type => bool) => nat) (hollight.UNIV::'A::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6948
      ((NUMERAL_BIT1::nat => nat) (0::nat)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6949
  by (import hollight DEF_dimindex)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6950
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6951
lemma HAS_SIZE_FINITE_IMAGE: "(All::(('A::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6952
 (%s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6953
     (HAS_SIZE::('A::type finite_image => bool) => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6954
      (hollight.UNIV::'A::type finite_image => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6955
      ((dimindex::('A::type => bool) => nat) s))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6956
  by (import hollight HAS_SIZE_FINITE_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6957
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6958
lemma CARD_FINITE_IMAGE: "(All::(('A::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6959
 (%s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6960
     (op =::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6961
      ((CARD::('A::type finite_image => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6962
        (hollight.UNIV::'A::type finite_image => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6963
      ((dimindex::('A::type => bool) => nat) s))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6964
  by (import hollight CARD_FINITE_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6965
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6966
lemma FINITE_FINITE_IMAGE: "(FINITE::('A::type finite_image => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6967
 (hollight.UNIV::'A::type finite_image => bool)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6968
  by (import hollight FINITE_FINITE_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6969
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6970
lemma DIMINDEX_NONZERO: "ALL s::'A::type => bool. dimindex s ~= 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6971
  by (import hollight DIMINDEX_NONZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6972
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  6973
lemma DIMINDEX_GE_1: "ALL x::'A::type => bool. <= (NUMERAL_BIT1 0) (dimindex x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6974
  by (import hollight DIMINDEX_GE_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6975
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6976
lemma DIMINDEX_FINITE_IMAGE: "ALL (s::'A::type finite_image => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6977
   dimindex s = dimindex t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6978
  by (import hollight DIMINDEX_FINITE_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6979
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  6980
definition finite_index :: "nat => 'A" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6981
  "(op ==::(nat => 'A::type) => (nat => 'A::type) => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6982
 (finite_index::nat => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6983
 ((Eps::((nat => 'A::type) => bool) => nat => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6984
   (%f::nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6985
       (All::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6986
        (%x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6987
            (Ex1::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6988
             (%n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6989
                 (op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6990
                  ((<=::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6991
                    ((NUMERAL_BIT1::nat => nat) (0::nat)) n)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6992
                  ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6993
                    ((<=::nat => nat => bool) n
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6994
                      ((dimindex::('A::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6995
                        (hollight.UNIV::'A::type => bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6996
                    ((op =::'A::type => 'A::type => bool) (f n) x))))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6997
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6998
lemma DEF_finite_index: "(op =::(nat => 'A::type) => (nat => 'A::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  6999
 (finite_index::nat => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7000
 ((Eps::((nat => 'A::type) => bool) => nat => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7001
   (%f::nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7002
       (All::('A::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7003
        (%x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7004
            (Ex1::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7005
             (%n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7006
                 (op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7007
                  ((<=::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7008
                    ((NUMERAL_BIT1::nat => nat) (0::nat)) n)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7009
                  ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7010
                    ((<=::nat => nat => bool) n
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7011
                      ((dimindex::('A::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7012
                        (hollight.UNIV::'A::type => bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7013
                    ((op =::'A::type => 'A::type => bool) (f n) x))))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7014
  by (import hollight DEF_finite_index)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7015
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7016
lemma FINITE_INDEX_WORKS_FINITE: "(op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7017
 ((FINITE::('N::type => bool) => bool) (hollight.UNIV::'N::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7018
 ((All::('N::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7019
   (%x::'N::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7020
       (Ex1::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7021
        (%xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7022
            (op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7023
             ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7024
               xa)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7025
             ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7026
               ((<=::nat => nat => bool) xa
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7027
                 ((dimindex::('N::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7028
                   (hollight.UNIV::'N::type => bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7029
               ((op =::'N::type => 'N::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7030
                 ((finite_index::nat => 'N::type) xa) x)))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7031
  by (import hollight FINITE_INDEX_WORKS_FINITE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7032
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7033
lemma FINITE_INDEX_WORKS: "(All::('A::type finite_image => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7034
 (%i::'A::type finite_image.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7035
     (Ex1::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7036
      (%n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7037
          (op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7038
           ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7039
             n)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7040
           ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7041
             ((<=::nat => nat => bool) n
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7042
               ((dimindex::('A::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7043
                 (hollight.UNIV::'A::type => bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7044
             ((op =::'A::type finite_image => 'A::type finite_image => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7045
               ((finite_index::nat => 'A::type finite_image) n) i))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7046
  by (import hollight FINITE_INDEX_WORKS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7047
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7048
lemma FINITE_INDEX_INJ: "(All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7049
 (%x::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7050
     (All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7051
      (%xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7052
          (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7053
           ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7054
             ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7055
               x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7056
             ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7057
               ((<=::nat => nat => bool) x
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7058
                 ((dimindex::('A::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7059
                   (hollight.UNIV::'A::type => bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7060
               ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7061
                 ((<=::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7062
                   ((NUMERAL_BIT1::nat => nat) (0::nat)) xa)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7063
                 ((<=::nat => nat => bool) xa
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7064
                   ((dimindex::('A::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7065
                     (hollight.UNIV::'A::type => bool))))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7066
           ((op =::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7067
             ((op =::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7068
               ((finite_index::nat => 'A::type) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7069
               ((finite_index::nat => 'A::type) xa))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7070
             ((op =::nat => nat => bool) x xa))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7071
  by (import hollight FINITE_INDEX_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7072
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7073
lemma FORALL_FINITE_INDEX: "(op =::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7074
 ((All::('N::type finite_image => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7075
   (P::'N::type finite_image => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7076
 ((All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7077
   (%i::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7078
       (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7079
        ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7080
          ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7081
          ((<=::nat => nat => bool) i
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7082
            ((dimindex::('N::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7083
              (hollight.UNIV::'N::type => bool))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7084
        (P ((finite_index::nat => 'N::type finite_image) i))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7085
  by (import hollight FORALL_FINITE_INDEX)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7086
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7087
typedef (open) ('A, 'B) cart = "(Collect::(('B::type finite_image => 'A::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7088
          => ('B::type finite_image => 'A::type) set)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7089
 (%f::'B::type finite_image => 'A::type. True::bool)"  morphisms "dest_cart" "mk_cart"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7090
  apply (rule light_ex_imp_nonempty[where t="(Eps::(('B::type finite_image => 'A::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7091
      => 'B::type finite_image => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7092
 (%f::'B::type finite_image => 'A::type. True::bool)"])
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7093
  by (import hollight TYDEF_cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7094
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7095
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7096
  dest_cart :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7097
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7098
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7099
  mk_cart :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7100
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7101
lemmas "TYDEF_cart_@intern" = typedef_hol2hollight 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7102
  [where a="a :: ('A, 'B) cart" and r=r ,
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7103
   OF type_definition_cart]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7104
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7105
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7106
  "$" :: "('q_46627, 'q_46634) cart => nat => 'q_46627" ("$")
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7107
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7108
defs
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7109
  "$_def": "$ ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7110
%(u::('q_46627::type, 'q_46634::type) cart) ua::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7111
   dest_cart u (finite_index ua)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7112
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7113
lemma "DEF_$": "$ =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7114
(%(u::('q_46627::type, 'q_46634::type) cart) ua::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7115
    dest_cart u (finite_index ua))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7116
  by (import hollight "DEF_$")
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7117
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7118
lemma CART_EQ: "(All::(('A::type, 'B::type) cart => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7119
 (%x::('A::type, 'B::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7120
     (All::(('A::type, 'B::type) cart => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7121
      (%y::('A::type, 'B::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7122
          (op =::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7123
           ((op =::('A::type, 'B::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7124
                   => ('A::type, 'B::type) cart => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7125
             x y)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7126
           ((All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7127
             (%xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7128
                 (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7129
                  ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7130
                    ((<=::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7131
                      ((NUMERAL_BIT1::nat => nat) (0::nat)) xa)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7132
                    ((<=::nat => nat => bool) xa
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7133
                      ((dimindex::('B::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7134
                        (hollight.UNIV::'B::type => bool))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7135
                  ((op =::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7136
                    (($::('A::type, 'B::type) cart => nat => 'A::type) x xa)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7137
                    (($::('A::type, 'B::type) cart => nat => 'A::type) y
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7138
                      xa))))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7139
  by (import hollight CART_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7140
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7141
definition lambda :: "(nat => 'A) => ('A, 'B) cart" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7142
  "(op ==::((nat => 'A::type) => ('A::type, 'B::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7143
        => ((nat => 'A::type) => ('A::type, 'B::type) cart) => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7144
 (lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7145
 (%u::nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7146
     (Eps::(('A::type, 'B::type) cart => bool) => ('A::type, 'B::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7147
      (%f::('A::type, 'B::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7148
          (All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7149
           (%i::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7150
               (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7151
                ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7152
                  ((<=::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7153
                    ((NUMERAL_BIT1::nat => nat) (0::nat)) i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7154
                  ((<=::nat => nat => bool) i
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7155
                    ((dimindex::('B::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7156
                      (hollight.UNIV::'B::type => bool))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7157
                ((op =::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7158
                  (($::('A::type, 'B::type) cart => nat => 'A::type) f i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7159
                  (u i)))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7160
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7161
lemma DEF_lambda: "(op =::((nat => 'A::type) => ('A::type, 'B::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7162
       => ((nat => 'A::type) => ('A::type, 'B::type) cart) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7163
 (lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7164
 (%u::nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7165
     (Eps::(('A::type, 'B::type) cart => bool) => ('A::type, 'B::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7166
      (%f::('A::type, 'B::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7167
          (All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7168
           (%i::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7169
               (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7170
                ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7171
                  ((<=::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7172
                    ((NUMERAL_BIT1::nat => nat) (0::nat)) i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7173
                  ((<=::nat => nat => bool) i
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7174
                    ((dimindex::('B::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7175
                      (hollight.UNIV::'B::type => bool))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7176
                ((op =::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7177
                  (($::('A::type, 'B::type) cart => nat => 'A::type) f i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7178
                  (u i)))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7179
  by (import hollight DEF_lambda)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7180
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7181
lemma LAMBDA_BETA: "(All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7182
 (%x::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7183
     (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7184
      ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7185
        ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7186
        ((<=::nat => nat => bool) x
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7187
          ((dimindex::('B::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7188
            (hollight.UNIV::'B::type => bool))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7189
      ((op =::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7190
        (($::('A::type, 'B::type) cart => nat => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7191
          ((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7192
            (g::nat => 'A::type))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7193
          x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7194
        (g x)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7195
  by (import hollight LAMBDA_BETA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7196
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7197
lemma LAMBDA_UNIQUE: "(All::(('A::type, 'B::type) cart => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7198
 (%x::('A::type, 'B::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7199
     (All::((nat => 'A::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7200
      (%xa::nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7201
          (op =::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7202
           ((All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7203
             (%i::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7204
                 (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7205
                  ((op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7206
                    ((<=::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7207
                      ((NUMERAL_BIT1::nat => nat) (0::nat)) i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7208
                    ((<=::nat => nat => bool) i
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7209
                      ((dimindex::('B::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7210
                        (hollight.UNIV::'B::type => bool))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7211
                  ((op =::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7212
                    (($::('A::type, 'B::type) cart => nat => 'A::type) x i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7213
                    (xa i))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7214
           ((op =::('A::type, 'B::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7215
                   => ('A::type, 'B::type) cart => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7216
             ((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) xa)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7217
             x)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7218
  by (import hollight LAMBDA_UNIQUE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7219
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7220
lemma LAMBDA_ETA: "ALL x::('q_46825::type, 'q_46829::type) cart. lambda ($ x) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7221
  by (import hollight LAMBDA_ETA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7222
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7223
typedef (open) ('A, 'B) finite_sum = "(Collect::('A::type finite_image + 'B::type finite_image => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7224
          => ('A::type finite_image + 'B::type finite_image) set)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7225
 (%x::'A::type finite_image + 'B::type finite_image. True::bool)"  morphisms "dest_finite_sum" "mk_finite_sum"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7226
  apply (rule light_ex_imp_nonempty[where t="(Eps::('A::type finite_image + 'B::type finite_image => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7227
      => 'A::type finite_image + 'B::type finite_image)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7228
 (%x::'A::type finite_image + 'B::type finite_image. True::bool)"])
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7229
  by (import hollight TYDEF_finite_sum)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7230
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7231
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7232
  dest_finite_sum :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7233
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7234
syntax
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7235
  mk_finite_sum :: _ 
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7236
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7237
lemmas "TYDEF_finite_sum_@intern" = typedef_hol2hollight 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7238
  [where a="a :: ('A, 'B) finite_sum" and r=r ,
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7239
   OF type_definition_finite_sum]
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7240
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7241
definition pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7242
  "(op ==::(('A::type, 'M::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7243
         => ('A::type, 'N::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7244
            => ('A::type, ('M::type, 'N::type) finite_sum) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7245
        => (('A::type, 'M::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7246
            => ('A::type, 'N::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7247
               => ('A::type, ('M::type, 'N::type) finite_sum) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7248
           => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7249
 (pastecart::('A::type, 'M::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7250
             => ('A::type, 'N::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7251
                => ('A::type, ('M::type, 'N::type) finite_sum) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7252
 (%(u::('A::type, 'M::type) cart) ua::('A::type, 'N::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7253
     (lambda::(nat => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7254
              => ('A::type, ('M::type, 'N::type) finite_sum) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7255
      (%i::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7256
          (COND::bool => 'A::type => 'A::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7257
           ((<=::nat => nat => bool) i
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7258
             ((dimindex::('M::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7259
               (hollight.UNIV::'M::type => bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7260
           (($::('A::type, 'M::type) cart => nat => 'A::type) u i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7261
           (($::('A::type, 'N::type) cart => nat => 'A::type) ua
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7262
             ((op -::nat => nat => nat) i
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7263
               ((dimindex::('M::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7264
                 (hollight.UNIV::'M::type => bool))))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7265
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7266
lemma DEF_pastecart: "(op =::(('A::type, 'M::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7267
        => ('A::type, 'N::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7268
           => ('A::type, ('M::type, 'N::type) finite_sum) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7269
       => (('A::type, 'M::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7270
           => ('A::type, 'N::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7271
              => ('A::type, ('M::type, 'N::type) finite_sum) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7272
          => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7273
 (pastecart::('A::type, 'M::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7274
             => ('A::type, 'N::type) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7275
                => ('A::type, ('M::type, 'N::type) finite_sum) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7276
 (%(u::('A::type, 'M::type) cart) ua::('A::type, 'N::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7277
     (lambda::(nat => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7278
              => ('A::type, ('M::type, 'N::type) finite_sum) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7279
      (%i::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7280
          (COND::bool => 'A::type => 'A::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7281
           ((<=::nat => nat => bool) i
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7282
             ((dimindex::('M::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7283
               (hollight.UNIV::'M::type => bool)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7284
           (($::('A::type, 'M::type) cart => nat => 'A::type) u i)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7285
           (($::('A::type, 'N::type) cart => nat => 'A::type) ua
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7286
             ((op -::nat => nat => nat) i
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7287
               ((dimindex::('M::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7288
                 (hollight.UNIV::'M::type => bool))))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7289
  by (import hollight DEF_pastecart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7290
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7291
definition fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7292
  "fstcart ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7293
%u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7294
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7295
lemma DEF_fstcart: "fstcart =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7296
(%u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7297
  by (import hollight DEF_fstcart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7298
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7299
definition sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7300
  "(op ==::(('A::type, ('M::type, 'N::type) finite_sum) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7301
         => ('A::type, 'N::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7302
        => (('A::type, ('M::type, 'N::type) finite_sum) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7303
            => ('A::type, 'N::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7304
           => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7305
 (sndcart::('A::type, ('M::type, 'N::type) finite_sum) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7306
           => ('A::type, 'N::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7307
 (%u::('A::type, ('M::type, 'N::type) finite_sum) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7308
     (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7309
      (%i::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7310
          ($::('A::type, ('M::type, 'N::type) finite_sum) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7311
              => nat => 'A::type)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7312
           u ((plus::nat => nat => nat) i
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7313
               ((dimindex::('M::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7314
                 (hollight.UNIV::'M::type => bool)))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7315
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7316
lemma DEF_sndcart: "(op =::(('A::type, ('M::type, 'N::type) finite_sum) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7317
        => ('A::type, 'N::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7318
       => (('A::type, ('M::type, 'N::type) finite_sum) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7319
           => ('A::type, 'N::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7320
          => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7321
 (sndcart::('A::type, ('M::type, 'N::type) finite_sum) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7322
           => ('A::type, 'N::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7323
 (%u::('A::type, ('M::type, 'N::type) finite_sum) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7324
     (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7325
      (%i::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7326
          ($::('A::type, ('M::type, 'N::type) finite_sum) cart
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7327
              => nat => 'A::type)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7328
           u ((plus::nat => nat => nat) i
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7329
               ((dimindex::('M::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7330
                 (hollight.UNIV::'M::type => bool)))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7331
  by (import hollight DEF_sndcart)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7332
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7333
lemma DIMINDEX_HAS_SIZE_FINITE_SUM: "(HAS_SIZE::(('M::type, 'N::type) finite_sum => bool) => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7334
 (hollight.UNIV::('M::type, 'N::type) finite_sum => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7335
 ((plus::nat => nat => nat)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7336
   ((dimindex::('M::type => bool) => nat) (hollight.UNIV::'M::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7337
   ((dimindex::('N::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7338
     (hollight.UNIV::'N::type => bool)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7339
  by (import hollight DIMINDEX_HAS_SIZE_FINITE_SUM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7340
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7341
lemma DIMINDEX_FINITE_SUM: "(op =::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7342
 ((dimindex::(('M::type, 'N::type) finite_sum => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7343
   (hollight.UNIV::('M::type, 'N::type) finite_sum => bool))
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7344
 ((plus::nat => nat => nat)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7345
   ((dimindex::('M::type => bool) => nat) (hollight.UNIV::'M::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7346
   ((dimindex::('N::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7347
     (hollight.UNIV::'N::type => bool)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7348
  by (import hollight DIMINDEX_FINITE_SUM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7349
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7350
lemma FSTCART_PASTECART: "ALL (x::('A::type, 'M::type) cart) xa::('A::type, 'N::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7351
   fstcart (pastecart x xa) = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7352
  by (import hollight FSTCART_PASTECART)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7353
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7354
lemma SNDCART_PASTECART: "ALL (x::('A::type, 'M::type) cart) xa::('A::type, 'N::type) cart.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7355
   sndcart (pastecart x xa) = xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7356
  by (import hollight SNDCART_PASTECART)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7357
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7358
lemma PASTECART_FST_SND: "ALL x::('q_47149::type, ('q_47146::type, 'q_47144::type) finite_sum) cart.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7359
   pastecart (fstcart x) (sndcart x) = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7360
  by (import hollight PASTECART_FST_SND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7361
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7362
lemma PASTECART_EQ: "ALL (x::('q_47187::type, ('q_47177::type, 'q_47188::type) finite_sum) cart)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7363
   y::('q_47187::type, ('q_47177::type, 'q_47188::type) finite_sum) cart.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7364
   (x = y) = (fstcart x = fstcart y & sndcart x = sndcart y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7365
  by (import hollight PASTECART_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7366
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7367
lemma FORALL_PASTECART: "All (P::('q_47208::type, ('q_47209::type, 'q_47210::type) finite_sum) cart
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7368
        => bool) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7369
(ALL (x::('q_47208::type, 'q_47209::type) cart)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7370
    y::('q_47208::type, 'q_47210::type) cart. P (pastecart x y))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7371
  by (import hollight FORALL_PASTECART)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7372
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7373
lemma EXISTS_PASTECART: "Ex (P::('q_47230::type, ('q_47231::type, 'q_47232::type) finite_sum) cart
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7374
       => bool) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7375
(EX (x::('q_47230::type, 'q_47231::type) cart)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7376
    y::('q_47230::type, 'q_47232::type) cart. P (pastecart x y))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7377
  by (import hollight EXISTS_PASTECART)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7378
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7379
lemma SURJECTIVE_IFF_INJECTIVE_GEN: "ALL (s::'A::type => bool) (t::'B::type => bool) f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7380
   FINITE s & FINITE t & CARD s = CARD t & SUBSET (IMAGE f s) t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7381
   (ALL y::'B::type. IN y t --> (EX x::'A::type. IN x s & f x = y)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7382
   (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:
diff changeset
  7383
  by (import hollight SURJECTIVE_IFF_INJECTIVE_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7384
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7385
lemma SURJECTIVE_IFF_INJECTIVE: "ALL (x::'A::type => bool) xa::'A::type => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7386
   FINITE x & SUBSET (IMAGE xa x) x -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7387
   (ALL y::'A::type. IN y x --> (EX xb::'A::type. IN xb x & xa xb = y)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7388
   (ALL (xb::'A::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7389
       IN xb x & IN y x & xa xb = xa y --> xb = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7390
  by (import hollight SURJECTIVE_IFF_INJECTIVE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7391
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7392
lemma IMAGE_IMP_INJECTIVE_GEN: "ALL (s::'A::type => bool) (t::'B::type => bool) f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7393
   FINITE s & CARD s = CARD t & IMAGE f s = t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7394
   (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:
diff changeset
  7395
  by (import hollight IMAGE_IMP_INJECTIVE_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7396
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7397
lemma IMAGE_IMP_INJECTIVE: "ALL (s::'q_47557::type => bool) f::'q_47557::type => 'q_47557::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7398
   FINITE s & IMAGE f s = s -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7399
   (ALL (x::'q_47557::type) y::'q_47557::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7400
       IN x s & IN y s & f x = f y --> x = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7401
  by (import hollight IMAGE_IMP_INJECTIVE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7402
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7403
lemma CARD_LE_INJ: "ALL (x::'A::type => bool) xa::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7404
   FINITE x & FINITE xa & <= (CARD x) (CARD xa) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7405
   (EX f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7406
       SUBSET (IMAGE f x) xa &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7407
       (ALL (xa::'A::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7408
           IN xa x & IN y x & f xa = f y --> xa = y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7409
  by (import hollight CARD_LE_INJ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7410
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7411
lemma FORALL_IN_CLAUSES: "(ALL x::'q_47663::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7412
    (ALL xa::'q_47663::type. IN xa EMPTY --> x xa) = True) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7413
(ALL (x::'q_47703::type => bool) (xa::'q_47703::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7414
    xb::'q_47703::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7415
    (ALL xc::'q_47703::type. IN xc (INSERT xa xb) --> x xc) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7416
    (x xa & (ALL xa::'q_47703::type. IN xa xb --> x xa)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7417
  by (import hollight FORALL_IN_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7418
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7419
lemma EXISTS_IN_CLAUSES: "(ALL x::'q_47723::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7420
    (EX xa::'q_47723::type. IN xa EMPTY & x xa) = False) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7421
(ALL (x::'q_47763::type => bool) (xa::'q_47763::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7422
    xb::'q_47763::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7423
    (EX xc::'q_47763::type. IN xc (INSERT xa xb) & x xc) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7424
    (x xa | (EX xa::'q_47763::type. IN xa xb & x xa)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7425
  by (import hollight EXISTS_IN_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7426
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7427
lemma SURJECTIVE_ON_RIGHT_INVERSE: "ALL (x::'q_47819::type => 'q_47820::type) xa::'q_47820::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7428
   (ALL xb::'q_47820::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7429
       IN xb xa -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7430
       (EX xa::'q_47819::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7431
           IN xa (s::'q_47819::type => bool) & x xa = xb)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7432
   (EX g::'q_47820::type => 'q_47819::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7433
       ALL y::'q_47820::type. IN y xa --> IN (g y) s & x (g y) = y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7434
  by (import hollight SURJECTIVE_ON_RIGHT_INVERSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7435
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7436
lemma INJECTIVE_ON_LEFT_INVERSE: "ALL (x::'q_47913::type => 'q_47916::type) xa::'q_47913::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7437
   (ALL (xb::'q_47913::type) y::'q_47913::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7438
       IN xb xa & IN y xa & x xb = x y --> xb = y) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7439
   (EX xb::'q_47916::type => 'q_47913::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7440
       ALL xc::'q_47913::type. IN xc xa --> xb (x xc) = xc)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7441
  by (import hollight INJECTIVE_ON_LEFT_INVERSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7442
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7443
lemma SURJECTIVE_RIGHT_INVERSE: "(ALL y::'q_47941::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7444
    EX x::'q_47944::type. (f::'q_47944::type => 'q_47941::type) x = y) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7445
(EX g::'q_47941::type => 'q_47944::type. ALL y::'q_47941::type. f (g y) = y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7446
  by (import hollight SURJECTIVE_RIGHT_INVERSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7447
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7448
lemma INJECTIVE_LEFT_INVERSE: "(ALL (x::'q_47978::type) xa::'q_47978::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7449
    (f::'q_47978::type => 'q_47981::type) x = f xa --> x = xa) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7450
(EX g::'q_47981::type => 'q_47978::type. ALL x::'q_47978::type. g (f x) = x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7451
  by (import hollight INJECTIVE_LEFT_INVERSE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7452
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7453
lemma FUNCTION_FACTORS_RIGHT: "ALL (x::'q_48017::type => 'q_48018::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7454
   xa::'q_48005::type => 'q_48018::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7455
   (ALL xb::'q_48017::type. EX y::'q_48005::type. xa y = x xb) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7456
   (EX xb::'q_48017::type => 'q_48005::type. x = xa o xb)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7457
  by (import hollight FUNCTION_FACTORS_RIGHT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7458
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7459
lemma FUNCTION_FACTORS_LEFT: "ALL (x::'q_48090::type => 'q_48091::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7460
   xa::'q_48090::type => 'q_48070::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7461
   (ALL (xb::'q_48090::type) y::'q_48090::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7462
       xa xb = xa y --> x xb = x y) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7463
   (EX xb::'q_48070::type => 'q_48091::type. x = xb o xa)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7464
  by (import hollight FUNCTION_FACTORS_LEFT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7465
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7466
definition dotdot :: "nat => nat => nat => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7467
  "dotdot ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7468
%(u::nat) ua::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7469
   GSPEC (%ub::nat. EX x::nat. SETSPEC ub (<= u x & <= x ua) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7470
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7471
lemma DEF__dot__dot_: "dotdot =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7472
(%(u::nat) ua::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7473
    GSPEC (%ub::nat. EX x::nat. SETSPEC ub (<= u x & <= x ua) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7474
  by (import hollight DEF__dot__dot_)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7475
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7476
lemma FINITE_NUMSEG: "ALL (m::nat) n::nat. FINITE (dotdot m n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7477
  by (import hollight FINITE_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7478
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7479
lemma NUMSEG_COMBINE_R: "ALL (x::'q_48166::type) (p::nat) m::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7480
   <= m p & <= p (n::nat) -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7481
   hollight.UNION (dotdot m p) (dotdot (p + NUMERAL_BIT1 0) n) = dotdot m n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7482
  by (import hollight NUMSEG_COMBINE_R)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7483
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7484
lemma NUMSEG_COMBINE_L: "ALL (x::'q_48204::type) (p::nat) m::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7485
   <= m p & <= p (n::nat) -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7486
   hollight.UNION (dotdot m (p - NUMERAL_BIT1 0)) (dotdot p n) = dotdot m n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7487
  by (import hollight NUMSEG_COMBINE_L)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7488
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7489
lemma NUMSEG_LREC: "ALL (x::nat) xa::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7490
   <= x xa --> INSERT x (dotdot (x + NUMERAL_BIT1 0) xa) = dotdot x xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7491
  by (import hollight NUMSEG_LREC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7492
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7493
lemma NUMSEG_RREC: "ALL (x::nat) xa::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7494
   <= x xa --> INSERT xa (dotdot x (xa - NUMERAL_BIT1 0)) = dotdot x xa"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7495
  by (import hollight NUMSEG_RREC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7496
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7497
lemma NUMSEG_REC: "ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7498
   <= x (Suc xa) --> dotdot x (Suc xa) = INSERT (Suc xa) (dotdot x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7499
  by (import hollight NUMSEG_REC)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7500
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7501
lemma IN_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat. IN xb (dotdot x xa) = (<= x xb & <= xb xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7502
  by (import hollight IN_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7503
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7504
lemma NUMSEG_SING: "ALL x::nat. dotdot x x = INSERT x EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7505
  by (import hollight NUMSEG_SING)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7506
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7507
lemma NUMSEG_EMPTY: "ALL (x::nat) xa::nat. (dotdot x xa = EMPTY) = < xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7508
  by (import hollight NUMSEG_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7509
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7510
lemma CARD_NUMSEG_LEMMA: "ALL (m::nat) d::nat. CARD (dotdot m (m + d)) = d + NUMERAL_BIT1 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7511
  by (import hollight CARD_NUMSEG_LEMMA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7512
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7513
lemma CARD_NUMSEG: "ALL (m::nat) n::nat. CARD (dotdot m n) = n + NUMERAL_BIT1 0 - m"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7514
  by (import hollight CARD_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7515
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7516
lemma HAS_SIZE_NUMSEG: "ALL (x::nat) xa::nat. HAS_SIZE (dotdot x xa) (xa + NUMERAL_BIT1 0 - x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7517
  by (import hollight HAS_SIZE_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7518
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7519
lemma CARD_NUMSEG_1: "ALL x::nat. CARD (dotdot (NUMERAL_BIT1 0) x) = x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7520
  by (import hollight CARD_NUMSEG_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7521
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7522
lemma HAS_SIZE_NUMSEG_1: "ALL x::nat. HAS_SIZE (dotdot (NUMERAL_BIT1 0) x) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7523
  by (import hollight HAS_SIZE_NUMSEG_1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7524
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7525
lemma NUMSEG_CLAUSES: "(ALL m::nat. dotdot m 0 = COND (m = 0) (INSERT 0 EMPTY) EMPTY) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7526
(ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7527
    dotdot m (Suc n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7528
    COND (<= m (Suc n)) (INSERT (Suc n) (dotdot m n)) (dotdot m n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7529
  by (import hollight NUMSEG_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7530
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7531
lemma FINITE_INDEX_NUMSEG: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7532
   FINITE s =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7533
   (EX f::nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7534
       (ALL (i::nat) j::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7535
           IN i (dotdot (NUMERAL_BIT1 0) (CARD s)) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7536
           IN j (dotdot (NUMERAL_BIT1 0) (CARD s)) & f i = f j -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7537
           i = j) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7538
       s = IMAGE f (dotdot (NUMERAL_BIT1 0) (CARD s)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7539
  by (import hollight FINITE_INDEX_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7540
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7541
lemma FINITE_INDEX_NUMBERS: "ALL s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7542
   FINITE s =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7543
   (EX (k::nat => bool) f::nat => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7544
       (ALL (i::nat) j::nat. IN i k & IN j k & f i = f j --> i = j) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7545
       FINITE k & s = IMAGE f k)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7546
  by (import hollight FINITE_INDEX_NUMBERS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7547
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7548
lemma DISJOINT_NUMSEG: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7549
   DISJOINT (dotdot x xa) (dotdot xb xc) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7550
   (< xa xb | < xc x | < xa x | < xc xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7551
  by (import hollight DISJOINT_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7552
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7553
lemma NUMSEG_ADD_SPLIT: "ALL (x::nat) (xa::nat) xb::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7554
   <= x (xa + NUMERAL_BIT1 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7555
   dotdot x (xa + xb) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7556
   hollight.UNION (dotdot x xa) (dotdot (xa + NUMERAL_BIT1 0) (xa + xb))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7557
  by (import hollight NUMSEG_ADD_SPLIT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7558
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7559
lemma NUMSEG_OFFSET_IMAGE: "ALL (x::nat) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7560
   dotdot (x + xb) (xa + xb) = IMAGE (%i::nat. i + xb) (dotdot x xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7561
  by (import hollight NUMSEG_OFFSET_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7562
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7563
lemma SUBSET_NUMSEG: "ALL (m::nat) (n::nat) (p::nat) q::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7564
   SUBSET (dotdot m n) (dotdot p q) = (< n m | <= p m & <= n q)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7565
  by (import hollight SUBSET_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7566
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7567
definition neutral :: "('q_48985 => 'q_48985 => 'q_48985) => 'q_48985" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7568
  "neutral ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7569
%u::'q_48985::type => 'q_48985::type => 'q_48985::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7570
   SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7571
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7572
lemma DEF_neutral: "neutral =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7573
(%u::'q_48985::type => 'q_48985::type => 'q_48985::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7574
    SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7575
  by (import hollight DEF_neutral)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7576
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7577
definition monoidal :: "('A => 'A => 'A) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7578
  "monoidal ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7579
%u::'A::type => 'A::type => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7580
   (ALL (x::'A::type) y::'A::type. u x y = u y x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7581
   (ALL (x::'A::type) (y::'A::type) z::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7582
       u x (u y z) = u (u x y) z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7583
   (ALL x::'A::type. u (neutral u) x = x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7584
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7585
lemma DEF_monoidal: "monoidal =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7586
(%u::'A::type => 'A::type => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7587
    (ALL (x::'A::type) y::'A::type. u x y = u y x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7588
    (ALL (x::'A::type) (y::'A::type) z::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7589
        u x (u y z) = u (u x y) z) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7590
    (ALL x::'A::type. u (neutral u) x = x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7591
  by (import hollight DEF_monoidal)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7592
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7593
definition support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7594
  "support ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7595
%(u::'B::type => 'B::type => 'B::type) (ua::'A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7596
   ub::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7597
   GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7598
    (%uc::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7599
        EX x::'A::type. SETSPEC uc (IN x ub & ua x ~= neutral u) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7600
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7601
lemma DEF_support: "support =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7602
(%(u::'B::type => 'B::type => 'B::type) (ua::'A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7603
    ub::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7604
    GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7605
     (%uc::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7606
         EX x::'A::type. SETSPEC uc (IN x ub & ua x ~= neutral u) x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7607
  by (import hollight DEF_support)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7608
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7609
definition iterate :: "('q_49090 => 'q_49090 => 'q_49090)
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7610
=> ('A => bool) => ('A => 'q_49090) => 'q_49090" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7611
  "iterate ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7612
%(u::'q_49090::type => 'q_49090::type => 'q_49090::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7613
   (ua::'A::type => bool) ub::'A::type => 'q_49090::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7614
   ITSET (%x::'A::type. u (ub x)) (support u ub ua) (neutral u)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7615
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7616
lemma DEF_iterate: "iterate =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7617
(%(u::'q_49090::type => 'q_49090::type => 'q_49090::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7618
    (ua::'A::type => bool) ub::'A::type => 'q_49090::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7619
    ITSET (%x::'A::type. u (ub x)) (support u ub ua) (neutral u))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7620
  by (import hollight DEF_iterate)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7621
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7622
lemma IN_SUPPORT: "ALL (x::'q_49133::type => 'q_49133::type => 'q_49133::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7623
   (xa::'q_49136::type => 'q_49133::type) (xb::'q_49136::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7624
   xc::'q_49136::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7625
   IN xb (support x xa xc) = (IN xb xc & xa xb ~= neutral x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7626
  by (import hollight IN_SUPPORT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7627
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7628
lemma SUPPORT_SUPPORT: "ALL (x::'q_49158::type => 'q_49158::type => 'q_49158::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7629
   (xa::'q_49169::type => 'q_49158::type) xb::'q_49169::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7630
   support x xa (support x xa xb) = support x xa xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7631
  by (import hollight SUPPORT_SUPPORT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7632
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7633
lemma SUPPORT_EMPTY: "ALL (x::'q_49194::type => 'q_49194::type => 'q_49194::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7634
   (xa::'q_49208::type => 'q_49194::type) xb::'q_49208::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7635
   (ALL xc::'q_49208::type. IN xc xb --> xa xc = neutral x) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7636
   (support x xa xb = EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7637
  by (import hollight SUPPORT_EMPTY)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7638
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7639
lemma SUPPORT_SUBSET: "ALL (x::'q_49228::type => 'q_49228::type => 'q_49228::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7640
   (xa::'q_49229::type => 'q_49228::type) xb::'q_49229::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7641
   SUBSET (support x xa xb) xb"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7642
  by (import hollight SUPPORT_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7643
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7644
lemma FINITE_SUPPORT: "ALL (u::'q_49252::type => 'q_49252::type => 'q_49252::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7645
   (f::'q_49246::type => 'q_49252::type) s::'q_49246::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7646
   FINITE s --> FINITE (support u f s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7647
  by (import hollight FINITE_SUPPORT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7648
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7649
lemma SUPPORT_CLAUSES: "(ALL x::'q_49270::type => 'q_49300::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7650
    support (u_4247::'q_49300::type => 'q_49300::type => 'q_49300::type) x
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7651
     EMPTY =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7652
    EMPTY) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7653
(ALL (x::'q_49318::type => 'q_49300::type) (xa::'q_49318::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7654
    xb::'q_49318::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7655
    support u_4247 x (INSERT xa xb) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7656
    COND (x xa = neutral u_4247) (support u_4247 x xb)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7657
     (INSERT xa (support u_4247 x xb))) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7658
(ALL (x::'q_49351::type => 'q_49300::type) (xa::'q_49351::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7659
    xb::'q_49351::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7660
    support u_4247 x (DELETE xb xa) = DELETE (support u_4247 x xb) xa) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7661
(ALL (x::'q_49389::type => 'q_49300::type) (xa::'q_49389::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7662
    xb::'q_49389::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7663
    support u_4247 x (hollight.UNION xa xb) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7664
    hollight.UNION (support u_4247 x xa) (support u_4247 x xb)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7665
(ALL (x::'q_49427::type => 'q_49300::type) (xa::'q_49427::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7666
    xb::'q_49427::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7667
    support u_4247 x (hollight.INTER xa xb) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7668
    hollight.INTER (support u_4247 x xa) (support u_4247 x xb)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7669
(ALL (x::'q_49463::type => 'q_49300::type) (xa::'q_49463::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7670
    xb::'q_49463::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7671
    support u_4247 x (DIFF xa xb) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7672
    DIFF (support u_4247 x xa) (support u_4247 x xb))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7673
  by (import hollight SUPPORT_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7674
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7675
lemma ITERATE_SUPPORT: "ALL (x::'q_49484::type => 'q_49484::type => 'q_49484::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7676
   (xa::'q_49485::type => 'q_49484::type) xb::'q_49485::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7677
   FINITE (support x xa xb) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7678
   iterate x (support x xa xb) xa = iterate x xb xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7679
  by (import hollight ITERATE_SUPPORT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7680
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7681
lemma SUPPORT_DELTA: "ALL (x::'q_49529::type => 'q_49529::type => 'q_49529::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7682
   (xa::'q_49557::type => bool) (xb::'q_49557::type => 'q_49529::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7683
   xc::'q_49557::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7684
   support x (%xa::'q_49557::type. COND (xa = xc) (xb xa) (neutral x)) xa =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7685
   COND (IN xc xa) (support x xb (INSERT xc EMPTY)) EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7686
  by (import hollight SUPPORT_DELTA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7687
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7688
lemma FINITE_SUPPORT_DELTA: "ALL (x::'q_49578::type => 'q_49578::type => 'q_49578::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7689
   (xa::'q_49587::type => 'q_49578::type) xb::'q_49587::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7690
   FINITE
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7691
    (support x (%xc::'q_49587::type. COND (xc = xb) (xa xc) (neutral x))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7692
      (s::'q_49587::type => bool))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7693
  by (import hollight FINITE_SUPPORT_DELTA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7694
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7695
lemma ITERATE_CLAUSES_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7696
   monoidal u_4247 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7697
   (ALL f::'A::type => 'B::type. iterate u_4247 EMPTY f = neutral u_4247) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7698
   (ALL (f::'A::type => 'B::type) (x::'A::type) s::'A::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7699
       monoidal u_4247 & FINITE (support u_4247 f s) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7700
       iterate u_4247 (INSERT x s) f =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7701
       COND (IN x s) (iterate u_4247 s f)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7702
        (u_4247 (f x) (iterate u_4247 s f)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7703
  by (import hollight ITERATE_CLAUSES_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7704
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7705
lemma ITERATE_CLAUSES: "ALL x::'q_49755::type => 'q_49755::type => 'q_49755::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7706
   monoidal x -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7707
   (ALL f::'q_49713::type => 'q_49755::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7708
       iterate x EMPTY f = neutral x) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7709
   (ALL (f::'q_49757::type => 'q_49755::type) (xa::'q_49757::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7710
       s::'q_49757::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7711
       FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7712
       iterate x (INSERT xa s) f =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7713
       COND (IN xa s) (iterate x s f) (x (f xa) (iterate x s f)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7714
  by (import hollight ITERATE_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7715
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7716
lemma ITERATE_UNION: "ALL u_4247::'q_49843::type => 'q_49843::type => 'q_49843::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7717
   monoidal u_4247 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7718
   (ALL (f::'q_49828::type => 'q_49843::type) (s::'q_49828::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7719
       x::'q_49828::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7720
       FINITE s & FINITE x & DISJOINT s x -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7721
       iterate u_4247 (hollight.UNION s x) f =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7722
       u_4247 (iterate u_4247 s f) (iterate u_4247 x f))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7723
  by (import hollight ITERATE_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7724
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7725
lemma ITERATE_UNION_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7726
   monoidal u_4247 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7727
   (ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'A::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7728
       FINITE (support u_4247 f s) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7729
       FINITE (support u_4247 f t) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7730
       DISJOINT (support u_4247 f s) (support u_4247 f t) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7731
       iterate u_4247 (hollight.UNION s t) f =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7732
       u_4247 (iterate u_4247 s f) (iterate u_4247 t f))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7733
  by (import hollight ITERATE_UNION_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7734
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7735
lemma ITERATE_DIFF: "ALL u::'q_49986::type => 'q_49986::type => 'q_49986::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7736
   monoidal u -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7737
   (ALL (f::'q_49982::type => 'q_49986::type) (s::'q_49982::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7738
       t::'q_49982::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7739
       FINITE s & SUBSET t s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7740
       u (iterate u (DIFF s t) f) (iterate u t f) = iterate u s f)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7741
  by (import hollight ITERATE_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7742
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7743
lemma ITERATE_DIFF_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7744
   monoidal u_4247 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7745
   (ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'A::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7746
       FINITE (support u_4247 f s) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7747
       SUBSET (support u_4247 f t) (support u_4247 f s) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7748
       u_4247 (iterate u_4247 (DIFF s t) f) (iterate u_4247 t f) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7749
       iterate u_4247 s f)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7750
  by (import hollight ITERATE_DIFF_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7751
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7752
lemma ITERATE_CLOSED: "ALL u_4247::'B::type => 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7753
   monoidal u_4247 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7754
   (ALL P::'B::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7755
       P (neutral u_4247) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7756
       (ALL (x::'B::type) y::'B::type. P x & P y --> P (u_4247 x y)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7757
       (ALL (f::'A::type => 'B::type) x::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7758
           FINITE x & (ALL xa::'A::type. IN xa x --> P (f xa)) -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7759
           P (iterate u_4247 x f)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7760
  by (import hollight ITERATE_CLOSED)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7761
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7762
lemma ITERATE_CLOSED_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7763
   monoidal u_4247 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7764
   (ALL P::'B::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7765
       P (neutral u_4247) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7766
       (ALL (x::'B::type) y::'B::type. P x & P y --> P (u_4247 x y)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7767
       (ALL (f::'A::type => 'B::type) s::'A::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7768
           FINITE (support u_4247 f s) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7769
           (ALL x::'A::type. IN x s & f x ~= neutral u_4247 --> P (f x)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7770
           P (iterate u_4247 s f)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7771
  by (import hollight ITERATE_CLOSED_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7772
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7773
lemma ITERATE_RELATED: "ALL u_4247::'B::type => 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7774
   monoidal u_4247 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7775
   (ALL R::'B::type => 'B::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7776
       R (neutral u_4247) (neutral u_4247) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7777
       (ALL (x1::'B::type) (y1::'B::type) (x2::'B::type) y2::'B::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7778
           R x1 x2 & R y1 y2 --> R (u_4247 x1 y1) (u_4247 x2 y2)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7779
       (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7780
           x::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7781
           FINITE x & (ALL xa::'A::type. IN xa x --> R (f xa) (g xa)) -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7782
           R (iterate u_4247 x f) (iterate u_4247 x g)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7783
  by (import hollight ITERATE_RELATED)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7784
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7785
lemma ITERATE_EQ_NEUTRAL: "ALL u_4247::'B::type => 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7786
   monoidal u_4247 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7787
   (ALL (f::'A::type => 'B::type) s::'A::type => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7788
       (ALL x::'A::type. IN x s --> f x = neutral u_4247) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7789
       iterate u_4247 s f = neutral u_4247)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7790
  by (import hollight ITERATE_EQ_NEUTRAL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7791
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7792
lemma ITERATE_SING: "ALL x::'B::type => 'B::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7793
   monoidal x -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7794
   (ALL (f::'A::type => 'B::type) xa::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7795
       iterate x (INSERT xa EMPTY) f = f xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7796
  by (import hollight ITERATE_SING)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7797
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7798
lemma ITERATE_DELTA: "ALL u_4247::'q_50438::type => 'q_50438::type => 'q_50438::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7799
   monoidal u_4247 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7800
   (ALL (x::'q_50457::type => 'q_50438::type) (xa::'q_50457::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7801
       xb::'q_50457::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7802
       iterate u_4247 xb
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7803
        (%xb::'q_50457::type. COND (xb = xa) (x xb) (neutral u_4247)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7804
       COND (IN xa xb) (x xa) (neutral u_4247))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7805
  by (import hollight ITERATE_DELTA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7806
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7807
lemma ITERATE_IMAGE: "ALL u_4247::'q_50536::type => 'q_50536::type => 'q_50536::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7808
   monoidal u_4247 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7809
   (ALL (f::'q_50535::type => 'q_50507::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7810
       (g::'q_50507::type => 'q_50536::type) x::'q_50535::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7811
       FINITE x &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7812
       (ALL (xa::'q_50535::type) y::'q_50535::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7813
           IN xa x & IN y x & f xa = f y --> xa = y) -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7814
       iterate u_4247 (IMAGE f x) g = iterate u_4247 x (g o f))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7815
  by (import hollight ITERATE_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7816
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7817
lemma ITERATE_BIJECTION: "ALL u_4247::'B::type => 'B::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7818
   monoidal u_4247 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7819
   (ALL (f::'A::type => 'B::type) (p::'A::type => 'A::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7820
       s::'A::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7821
       FINITE s &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7822
       (ALL x::'A::type. IN x s --> IN (p x) s) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7823
       (ALL y::'A::type. IN y s --> (EX! x::'A::type. IN x s & p x = y)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7824
       iterate u_4247 s f = iterate u_4247 s (f o p))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7825
  by (import hollight ITERATE_BIJECTION)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7826
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7827
lemma ITERATE_ITERATE_PRODUCT: "ALL u_4247::'C::type => 'C::type => 'C::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7828
   monoidal u_4247 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7829
   (ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7830
       xb::'A::type => 'B::type => 'C::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7831
       FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7832
       iterate u_4247 x (%i::'A::type. iterate u_4247 (xa i) (xb i)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7833
       iterate u_4247
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7834
        (GSPEC
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7835
          (%u::'A::type * 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7836
              EX (i::'A::type) j::'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7837
                 SETSPEC u (IN i x & IN j (xa i)) (i, j)))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7838
        (GABS
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7839
          (%f::'A::type * 'B::type => 'C::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7840
              ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j))))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7841
  by (import hollight ITERATE_ITERATE_PRODUCT)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7842
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7843
lemma ITERATE_EQ: "ALL u_4247::'q_50867::type => 'q_50867::type => 'q_50867::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7844
   monoidal u_4247 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7845
   (ALL (f::'q_50871::type => 'q_50867::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7846
       (g::'q_50871::type => 'q_50867::type) x::'q_50871::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7847
       FINITE x & (ALL xa::'q_50871::type. IN xa x --> f xa = g xa) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7848
       iterate u_4247 x f = iterate u_4247 x g)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7849
  by (import hollight ITERATE_EQ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7850
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7851
lemma ITERATE_EQ_GENERAL: "ALL u_4247::'C::type => 'C::type => 'C::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7852
   monoidal u_4247 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7853
   (ALL (s::'A::type => bool) (t::'B::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7854
       (f::'A::type => 'C::type) (g::'B::type => 'C::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7855
       h::'A::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7856
       FINITE s &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7857
       (ALL y::'B::type. IN y t --> (EX! x::'A::type. IN x s & h x = y)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7858
       (ALL x::'A::type. IN x s --> IN (h x) t & g (h x) = f x) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7859
       iterate u_4247 s f = iterate u_4247 t g)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7860
  by (import hollight ITERATE_EQ_GENERAL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7861
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  7862
definition nsum :: "('q_51017 => bool) => ('q_51017 => nat) => nat" where 
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7863
  "(op ==::(('q_51017::type => bool) => ('q_51017::type => nat) => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7864
        => (('q_51017::type => bool) => ('q_51017::type => nat) => nat)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7865
           => prop)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7866
 (nsum::('q_51017::type => bool) => ('q_51017::type => nat) => nat)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7867
 ((iterate::(nat => nat => nat)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7868
            => ('q_51017::type => bool) => ('q_51017::type => nat) => nat)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7869
   (plus::nat => nat => nat))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7870
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7871
lemma DEF_nsum: "(op =::(('q_51017::type => bool) => ('q_51017::type => nat) => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7872
       => (('q_51017::type => bool) => ('q_51017::type => nat) => nat)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7873
          => bool)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7874
 (nsum::('q_51017::type => bool) => ('q_51017::type => nat) => nat)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7875
 ((iterate::(nat => nat => nat)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7876
            => ('q_51017::type => bool) => ('q_51017::type => nat) => nat)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7877
   (plus::nat => nat => nat))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7878
  by (import hollight DEF_nsum)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7879
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7880
lemma NEUTRAL_ADD: "(op =::nat => nat => bool)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7881
 ((neutral::(nat => nat => nat) => nat) (plus::nat => nat => nat)) (0::nat)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7882
  by (import hollight NEUTRAL_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7883
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7884
lemma NEUTRAL_MUL: "neutral op * = NUMERAL_BIT1 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7885
  by (import hollight NEUTRAL_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7886
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7887
lemma MONOIDAL_ADD: "(monoidal::(nat => nat => nat) => bool) (plus::nat => nat => nat)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7888
  by (import hollight MONOIDAL_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7889
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7890
lemma MONOIDAL_MUL: "(monoidal::(nat => nat => nat) => bool) (op *::nat => nat => nat)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7891
  by (import hollight MONOIDAL_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7892
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7893
lemma NSUM_CLAUSES: "(ALL x::'q_51055::type => nat. nsum EMPTY x = 0) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7894
(ALL (x::'q_51094::type) (xa::'q_51094::type => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7895
    xb::'q_51094::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7896
    FINITE xb -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7897
    nsum (INSERT x xb) xa = COND (IN x xb) (nsum xb xa) (xa x + nsum xb xa))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7898
  by (import hollight NSUM_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7899
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7900
lemma NSUM_UNION: "ALL (x::'q_51120::type => nat) (xa::'q_51120::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7901
   xb::'q_51120::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7902
   FINITE xa & FINITE xb & DISJOINT xa xb -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7903
   nsum (hollight.UNION xa xb) x = nsum xa x + nsum xb x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7904
  by (import hollight NSUM_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7905
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7906
lemma NSUM_DIFF: "ALL (f::'q_51175::type => nat) (s::'q_51175::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7907
   t::'q_51175::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7908
   FINITE s & SUBSET t s --> nsum (DIFF s t) f = nsum s f - nsum t f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7909
  by (import hollight NSUM_DIFF)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7910
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7911
lemma NSUM_SUPPORT: "ALL (x::'q_51214::type => nat) xa::'q_51214::type => bool.
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 19233
diff changeset
  7912
   FINITE (support plus x xa) --> nsum (support plus x xa) x = nsum xa x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7913
  by (import hollight NSUM_SUPPORT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7914
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7915
lemma NSUM_ADD: "ALL (f::'q_51260::type => nat) (g::'q_51260::type => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7916
   s::'q_51260::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7917
   FINITE s --> nsum s (%x::'q_51260::type. f x + g x) = nsum s f + nsum s g"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7918
  by (import hollight NSUM_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7919
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7920
lemma NSUM_CMUL: "ALL (f::'q_51298::type => nat) (c::nat) s::'q_51298::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7921
   FINITE s --> nsum s (%x::'q_51298::type. c * f x) = c * nsum s f"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7922
  by (import hollight NSUM_CMUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7923
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7924
lemma NSUM_LE: "ALL (x::'q_51337::type => nat) (xa::'q_51337::type => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7925
   xb::'q_51337::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7926
   FINITE xb & (ALL xc::'q_51337::type. IN xc xb --> <= (x xc) (xa xc)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7927
   <= (nsum xb x) (nsum xb xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7928
  by (import hollight NSUM_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7929
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7930
lemma NSUM_LT: "ALL (f::'A::type => nat) (g::'A::type => nat) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7931
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7932
   (ALL x::'A::type. IN x s --> <= (f x) (g x)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7933
   (EX x::'A::type. IN x s & < (f x) (g x)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7934
   < (nsum s f) (nsum s g)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7935
  by (import hollight NSUM_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7936
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7937
lemma NSUM_LT_ALL: "ALL (f::'q_51459::type => nat) (g::'q_51459::type => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7938
   s::'q_51459::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7939
   FINITE s &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7940
   s ~= EMPTY & (ALL x::'q_51459::type. IN x s --> < (f x) (g x)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7941
   < (nsum s f) (nsum s g)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7942
  by (import hollight NSUM_LT_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7943
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7944
lemma NSUM_EQ: "ALL (x::'q_51501::type => nat) (xa::'q_51501::type => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7945
   xb::'q_51501::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7946
   FINITE xb & (ALL xc::'q_51501::type. IN xc xb --> x xc = xa xc) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7947
   nsum xb x = nsum xb xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7948
  by (import hollight NSUM_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7949
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7950
lemma NSUM_CONST: "ALL (c::nat) s::'q_51531::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7951
   FINITE s --> nsum s (%n::'q_51531::type. c) = CARD s * c"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7952
  by (import hollight NSUM_CONST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7953
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7954
lemma NSUM_EQ_0: "ALL (x::'A::type => nat) xa::'A::type => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7955
   (ALL xb::'A::type. IN xb xa --> x xb = 0) --> nsum xa x = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7956
  by (import hollight NSUM_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7957
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7958
lemma NSUM_0: "ALL x::'A::type => bool. nsum x (%n::'A::type. 0) = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7959
  by (import hollight NSUM_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7960
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7961
lemma NSUM_POS_LE: "ALL (x::'q_51610::type => nat) xa::'q_51610::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7962
   FINITE xa & (ALL xb::'q_51610::type. IN xb xa --> <= 0 (x xb)) -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7963
   <= 0 (nsum xa x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7964
  by (import hollight NSUM_POS_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7965
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7966
lemma NSUM_POS_BOUND: "ALL (f::'A::type => nat) (b::nat) x::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7967
   FINITE x &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7968
   (ALL xa::'A::type. IN xa x --> <= 0 (f xa)) & <= (nsum x f) b -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7969
   (ALL xa::'A::type. IN xa x --> <= (f xa) b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7970
  by (import hollight NSUM_POS_BOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7971
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7972
lemma NSUM_POS_EQ_0: "ALL (x::'q_51745::type => nat) xa::'q_51745::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7973
   FINITE xa &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7974
   (ALL xb::'q_51745::type. IN xb xa --> <= 0 (x xb)) & nsum xa x = 0 -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7975
   (ALL xb::'q_51745::type. IN xb xa --> x xb = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7976
  by (import hollight NSUM_POS_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7977
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7978
lemma NSUM_SING: "ALL (x::'q_51765::type => nat) xa::'q_51765::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7979
   nsum (INSERT xa EMPTY) x = x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7980
  by (import hollight NSUM_SING)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7981
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7982
lemma NSUM_DELTA: "ALL (x::'A::type => bool) xa::'A::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  7983
   nsum x (%x::'A::type. COND (x = xa) (b::nat) 0) = COND (IN xa x) b 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7984
  by (import hollight NSUM_DELTA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7985
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7986
lemma NSUM_SWAP: "ALL (f::'A::type => 'B::type => nat) (x::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7987
   xa::'B::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7988
   FINITE x & FINITE xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7989
   nsum x (%i::'A::type. nsum xa (f i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7990
   nsum xa (%j::'B::type. nsum x (%i::'A::type. f i j))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7991
  by (import hollight NSUM_SWAP)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7992
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7993
lemma NSUM_IMAGE: "ALL (x::'q_51905::type => 'q_51881::type) (xa::'q_51881::type => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7994
   xb::'q_51905::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7995
   FINITE xb &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  7996
   (ALL (xa::'q_51905::type) y::'q_51905::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7997
       IN xa xb & IN y xb & x xa = x y --> xa = y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7998
   nsum (IMAGE x xb) xa = nsum xb (xa o x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  7999
  by (import hollight NSUM_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8000
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8001
lemma NSUM_SUPERSET: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8002
   FINITE u &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8003
   SUBSET u v & (ALL x::'A::type. IN x v & ~ IN x u --> f x = 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8004
   nsum v f = nsum u f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8005
  by (import hollight NSUM_SUPERSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8006
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8007
lemma NSUM_UNION_RZERO: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8008
   FINITE u & (ALL x::'A::type. IN x v & ~ IN x u --> f x = 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8009
   nsum (hollight.UNION u v) f = nsum u f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8010
  by (import hollight NSUM_UNION_RZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8011
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8012
lemma NSUM_UNION_LZERO: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8013
   FINITE v & (ALL x::'A::type. IN x u & ~ IN x v --> f x = 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8014
   nsum (hollight.UNION u v) f = nsum v f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8015
  by (import hollight NSUM_UNION_LZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8016
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8017
lemma NSUM_RESTRICT: "ALL (f::'q_52126::type => nat) s::'q_52126::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8018
   FINITE s -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8019
   nsum s (%x::'q_52126::type. COND (IN x s) (f x) 0) = nsum s f"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8020
  by (import hollight NSUM_RESTRICT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8021
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8022
lemma NSUM_BOUND: "ALL (x::'A::type => bool) (xa::'A::type => nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8023
   FINITE x & (ALL xc::'A::type. IN xc x --> <= (xa xc) xb) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8024
   <= (nsum x xa) (CARD x * xb)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8025
  by (import hollight NSUM_BOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8026
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8027
lemma NSUM_BOUND_GEN: "ALL (x::'A::type => bool) (xa::'q_52201::type) b::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8028
   FINITE x &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8029
   x ~= EMPTY &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8030
   (ALL xa::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8031
       IN xa x --> <= ((f::'A::type => nat) xa) (DIV b (CARD x))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8032
   <= (nsum x f) b"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8033
  by (import hollight NSUM_BOUND_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8034
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8035
lemma NSUM_BOUND_LT: "ALL (s::'A::type => bool) (f::'A::type => nat) b::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8036
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8037
   (ALL x::'A::type. IN x s --> <= (f x) b) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8038
   (EX x::'A::type. IN x s & < (f x) b) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8039
   < (nsum s f) (CARD s * b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8040
  by (import hollight NSUM_BOUND_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8041
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8042
lemma NSUM_BOUND_LT_ALL: "ALL (s::'q_52344::type => bool) (f::'q_52344::type => nat) b::nat.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8043
   FINITE s & s ~= EMPTY & (ALL x::'q_52344::type. IN x s --> < (f x) b) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8044
   < (nsum s f) (CARD s * b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8045
  by (import hollight NSUM_BOUND_LT_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8046
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8047
lemma NSUM_BOUND_LT_GEN: "ALL (s::'A::type => bool) (t::'q_52386::type) b::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8048
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8049
   s ~= EMPTY &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8050
   (ALL x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8051
       IN x s --> < ((f::'A::type => nat) x) (DIV b (CARD s))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8052
   < (nsum s f) b"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8053
  by (import hollight NSUM_BOUND_LT_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8054
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8055
lemma NSUM_UNION_EQ: "ALL (s::'q_52445::type => bool) (t::'q_52445::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8056
   u::'q_52445::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8057
   FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8058
   nsum s (f::'q_52445::type => nat) + nsum t f = nsum u f"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8059
  by (import hollight NSUM_UNION_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8060
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8061
lemma NSUM_EQ_SUPERSET: "ALL (f::'A::type => nat) (s::'A::type => bool) t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8062
   FINITE t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8063
   SUBSET t s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8064
   (ALL x::'A::type. IN x t --> f x = (g::'A::type => nat) x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8065
   (ALL x::'A::type. IN x s & ~ IN x t --> f x = 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8066
   nsum s f = nsum t g"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8067
  by (import hollight NSUM_EQ_SUPERSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8068
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8069
lemma NSUM_RESTRICT_SET: "ALL (s::'A::type => bool) (f::'A::type => nat) r::'q_52556::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8070
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8071
   nsum
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8072
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8073
      (%u::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8074
          EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8075
    f =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8076
   nsum s (%x::'A::type. COND (P x) (f x) 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8077
  by (import hollight NSUM_RESTRICT_SET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8078
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8079
lemma NSUM_NSUM_RESTRICT: "ALL (R::'q_52685::type => 'q_52684::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8080
   (f::'q_52685::type => 'q_52684::type => nat) (s::'q_52685::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8081
   t::'q_52684::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8082
   FINITE s & FINITE t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8083
   nsum s
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8084
    (%x::'q_52685::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8085
        nsum
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8086
         (GSPEC
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8087
           (%u::'q_52684::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8088
               EX y::'q_52684::type. SETSPEC u (IN y t & R x y) y))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8089
         (f x)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8090
   nsum t
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8091
    (%y::'q_52684::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8092
        nsum
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8093
         (GSPEC
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8094
           (%u::'q_52685::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8095
               EX x::'q_52685::type. SETSPEC u (IN x s & R x y) x))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8096
         (%x::'q_52685::type. f x y))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8097
  by (import hollight NSUM_NSUM_RESTRICT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8098
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8099
lemma CARD_EQ_NSUM: "ALL x::'q_52704::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8100
   FINITE x --> CARD x = nsum x (%x::'q_52704::type. NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8101
  by (import hollight CARD_EQ_NSUM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8102
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8103
lemma NSUM_MULTICOUNT_GEN: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8104
   (t::'B::type => bool) k::'B::type => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8105
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8106
   FINITE t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8107
   (ALL j::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8108
       IN j t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8109
       CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8110
        (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8111
          (%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8112
       k j) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8113
   nsum s
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8114
    (%i::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8115
        CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8116
         (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8117
           (%u::'B::type. EX j::'B::type. SETSPEC u (IN j t & R i j) j))) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8118
   nsum t k"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8119
  by (import hollight NSUM_MULTICOUNT_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8120
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8121
lemma NSUM_MULTICOUNT: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8122
   (t::'B::type => bool) k::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8123
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8124
   FINITE t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8125
   (ALL j::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8126
       IN j t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8127
       CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8128
        (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8129
          (%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8130
       k) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8131
   nsum s
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8132
    (%i::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8133
        CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8134
         (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8135
           (%u::'B::type. EX j::'B::type. SETSPEC u (IN j t & R i j) j))) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8136
   k * CARD t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8137
  by (import hollight NSUM_MULTICOUNT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8138
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8139
lemma NSUM_IMAGE_GEN: "ALL (f::'A::type => 'B::type) (g::'A::type => nat) s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8140
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8141
   nsum s g =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8142
   nsum (IMAGE f s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8143
    (%y::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8144
        nsum
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8145
         (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8146
           (%u::'A::type. EX x::'A::type. SETSPEC u (IN x s & f x = y) x))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8147
         g)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8148
  by (import hollight NSUM_IMAGE_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8149
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8150
lemma NSUM_SUBSET: "ALL (u::'A::type => bool) (v::'A::type => bool) f::'A::type => nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8151
   FINITE u & FINITE v & (ALL x::'A::type. IN x (DIFF u v) --> f x = 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8152
   <= (nsum u f) (nsum v f)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8153
  by (import hollight NSUM_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8154
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8155
lemma NSUM_SUBSET_SIMPLE: "ALL (u::'q_53164::type => bool) (v::'q_53164::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8156
   f::'q_53164::type => nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8157
   FINITE v & SUBSET u v --> <= (nsum u f) (nsum v f)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8158
  by (import hollight NSUM_SUBSET_SIMPLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8159
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8160
lemma NSUM_ADD_NUMSEG: "ALL (x::nat => nat) (xa::nat => nat) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8161
   nsum (dotdot xb xc) (%i::nat. x i + xa i) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8162
   nsum (dotdot xb xc) x + nsum (dotdot xb xc) xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8163
  by (import hollight NSUM_ADD_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8164
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8165
lemma NSUM_CMUL_NUMSEG: "ALL (x::nat => nat) (xa::nat) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8166
   nsum (dotdot xb xc) (%i::nat. xa * x i) = xa * nsum (dotdot xb xc) x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8167
  by (import hollight NSUM_CMUL_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8168
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8169
lemma NSUM_LE_NUMSEG: "ALL (x::nat => nat) (xa::nat => nat) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8170
   (ALL i::nat. <= xb i & <= i xc --> <= (x i) (xa i)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8171
   <= (nsum (dotdot xb xc) x) (nsum (dotdot xb xc) xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8172
  by (import hollight NSUM_LE_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8173
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8174
lemma NSUM_EQ_NUMSEG: "ALL (f::nat => nat) (g::nat => nat) (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8175
   (ALL i::nat. <= m i & <= i n --> f i = g i) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8176
   nsum (dotdot m n) f = nsum (dotdot m n) g"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8177
  by (import hollight NSUM_EQ_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8178
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8179
lemma NSUM_CONST_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8180
   nsum (dotdot xa xb) (%n::nat. x) = (xb + NUMERAL_BIT1 0 - xa) * x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8181
  by (import hollight NSUM_CONST_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8182
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8183
lemma NSUM_EQ_0_NUMSEG: "ALL (x::nat => nat) xa::'q_53403::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8184
   (ALL i::nat. <= (m::nat) i & <= i (n::nat) --> x i = 0) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8185
   nsum (dotdot m n) x = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8186
  by (import hollight NSUM_EQ_0_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8187
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8188
lemma NSUM_TRIV_NUMSEG: "ALL (f::nat => nat) (m::nat) n::nat. < n m --> nsum (dotdot m n) f = 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8189
  by (import hollight NSUM_TRIV_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8190
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8191
lemma NSUM_POS_LE_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat => nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8192
   (ALL p::nat. <= x p & <= p xa --> <= 0 (xb p)) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8193
   <= 0 (nsum (dotdot x xa) xb)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8194
  by (import hollight NSUM_POS_LE_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8195
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8196
lemma NSUM_POS_EQ_0_NUMSEG: "ALL (f::nat => nat) (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8197
   (ALL p::nat. <= m p & <= p n --> <= 0 (f p)) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8198
   nsum (dotdot m n) f = 0 -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8199
   (ALL p::nat. <= m p & <= p n --> f p = 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8200
  by (import hollight NSUM_POS_EQ_0_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8201
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8202
lemma NSUM_SING_NUMSEG: "ALL (x::nat => nat) xa::nat. nsum (dotdot xa xa) x = x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8203
  by (import hollight NSUM_SING_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8204
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8205
lemma NSUM_CLAUSES_NUMSEG: "(ALL x::nat. nsum (dotdot x 0) (f::nat => nat) = COND (x = 0) (f 0) 0) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8206
(ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8207
    nsum (dotdot x (Suc xa)) f =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8208
    COND (<= x (Suc xa)) (nsum (dotdot x xa) f + f (Suc xa))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8209
     (nsum (dotdot x xa) f))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8210
  by (import hollight NSUM_CLAUSES_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8211
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8212
lemma NSUM_SWAP_NUMSEG: "ALL (a::nat) (b::nat) (c::nat) (d::nat) f::nat => nat => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8213
   nsum (dotdot a b) (%i::nat. nsum (dotdot c d) (f i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8214
   nsum (dotdot c d) (%j::nat. nsum (dotdot a b) (%i::nat. f i j))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8215
  by (import hollight NSUM_SWAP_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8216
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8217
lemma NSUM_ADD_SPLIT: "ALL (x::nat => nat) (xa::nat) (xb::nat) xc::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8218
   <= xa (xb + NUMERAL_BIT1 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8219
   nsum (dotdot xa (xb + xc)) x =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8220
   nsum (dotdot xa xb) x + nsum (dotdot (xb + NUMERAL_BIT1 0) (xb + xc)) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8221
  by (import hollight NSUM_ADD_SPLIT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8222
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8223
lemma NSUM_OFFSET: "ALL (x::nat => nat) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8224
   nsum (dotdot (xa + xb) ((n::nat) + xb)) x =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8225
   nsum (dotdot xa n) (%i::nat. x (i + xb))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8226
  by (import hollight NSUM_OFFSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8227
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8228
lemma NSUM_OFFSET_0: "ALL (x::nat => nat) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8229
   <= xa xb -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8230
   nsum (dotdot xa xb) x = nsum (dotdot 0 (xb - xa)) (%i::nat. x (i + xa))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8231
  by (import hollight NSUM_OFFSET_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8232
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8233
lemma NSUM_CLAUSES_LEFT: "ALL (x::nat => nat) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8234
   <= xa xb -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8235
   nsum (dotdot xa xb) x = x xa + nsum (dotdot (xa + NUMERAL_BIT1 0) xb) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8236
  by (import hollight NSUM_CLAUSES_LEFT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8237
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8238
lemma NSUM_CLAUSES_RIGHT: "ALL (f::nat => nat) (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8239
   < 0 n & <= m n -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8240
   nsum (dotdot m n) f = nsum (dotdot m (n - NUMERAL_BIT1 0)) f + f n"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8241
  by (import hollight NSUM_CLAUSES_RIGHT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8242
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8243
lemma NSUM_BIJECTION: "ALL (x::'A::type => nat) (xa::'A::type => 'A::type) xb::'A::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8244
   FINITE xb &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8245
   (ALL x::'A::type. IN x xb --> IN (xa x) xb) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8246
   (ALL y::'A::type. IN y xb --> (EX! x::'A::type. IN x xb & xa x = y)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8247
   nsum xb x = nsum xb (x o xa)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8248
  by (import hollight NSUM_BIJECTION)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8249
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8250
lemma NSUM_NSUM_PRODUCT: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8251
   xb::'A::type => 'B::type => nat.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8252
   FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8253
   nsum x (%x::'A::type. nsum (xa x) (xb x)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8254
   nsum
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8255
    (GSPEC
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8256
      (%u::'A::type * 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8257
          EX (i::'A::type) j::'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8258
             SETSPEC u (IN i x & IN j (xa i)) (i, j)))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8259
    (GABS
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8260
      (%f::'A::type * 'B::type => nat.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8261
          ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j)))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8262
  by (import hollight NSUM_NSUM_PRODUCT)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8263
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8264
lemma NSUM_EQ_GENERAL: "ALL (x::'A::type => bool) (xa::'B::type => bool) (xb::'A::type => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8265
   (xc::'B::type => nat) xd::'A::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8266
   FINITE x &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8267
   (ALL y::'B::type. IN y xa --> (EX! xa::'A::type. IN xa x & xd xa = y)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8268
   (ALL xe::'A::type. IN xe x --> IN (xd xe) xa & xc (xd xe) = xb xe) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8269
   nsum x xb = nsum xa xc"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8270
  by (import hollight NSUM_EQ_GENERAL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8271
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8272
consts
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8273
  sum :: "('q_54215 => bool) => ('q_54215 => hollight.real) => hollight.real" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8274
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8275
defs
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8276
  sum_def: "(op ==::(('q_54215::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8277
         => ('q_54215::type => hollight.real) => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8278
        => (('q_54215::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8279
            => ('q_54215::type => hollight.real) => hollight.real)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8280
           => prop)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8281
 (hollight.sum::('q_54215::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8282
                => ('q_54215::type => hollight.real) => hollight.real)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8283
 ((iterate::(hollight.real => hollight.real => hollight.real)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8284
            => ('q_54215::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8285
               => ('q_54215::type => hollight.real) => hollight.real)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8286
   (real_add::hollight.real => hollight.real => hollight.real))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8287
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8288
lemma DEF_sum: "(op =::(('q_54215::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8289
        => ('q_54215::type => hollight.real) => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8290
       => (('q_54215::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8291
           => ('q_54215::type => hollight.real) => hollight.real)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8292
          => bool)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8293
 (hollight.sum::('q_54215::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8294
                => ('q_54215::type => hollight.real) => hollight.real)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8295
 ((iterate::(hollight.real => hollight.real => hollight.real)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8296
            => ('q_54215::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8297
               => ('q_54215::type => hollight.real) => hollight.real)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8298
   (real_add::hollight.real => hollight.real => hollight.real))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8299
  by (import hollight DEF_sum)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8300
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8301
lemma NEUTRAL_REAL_ADD: "neutral real_add = real_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8302
  by (import hollight NEUTRAL_REAL_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8303
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8304
lemma NEUTRAL_REAL_MUL: "neutral real_mul = real_of_num (NUMERAL_BIT1 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8305
  by (import hollight NEUTRAL_REAL_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8306
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8307
lemma MONOIDAL_REAL_ADD: "monoidal real_add"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8308
  by (import hollight MONOIDAL_REAL_ADD)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8309
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8310
lemma MONOIDAL_REAL_MUL: "monoidal real_mul"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8311
  by (import hollight MONOIDAL_REAL_MUL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8312
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8313
lemma SUM_CLAUSES: "(ALL x::'q_54257::type => hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8314
    hollight.sum EMPTY x = real_of_num 0) &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8315
(ALL (x::'q_54298::type) (xa::'q_54298::type => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8316
    xb::'q_54298::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8317
    FINITE xb -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8318
    hollight.sum (INSERT x xb) xa =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8319
    COND (IN x xb) (hollight.sum xb xa)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8320
     (real_add (xa x) (hollight.sum xb xa)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8321
  by (import hollight SUM_CLAUSES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8322
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8323
lemma SUM_UNION: "ALL (x::'q_54324::type => hollight.real) (xa::'q_54324::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8324
   xb::'q_54324::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8325
   FINITE xa & FINITE xb & DISJOINT xa xb -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8326
   hollight.sum (hollight.UNION xa xb) x =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8327
   real_add (hollight.sum xa x) (hollight.sum xb x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8328
  by (import hollight SUM_UNION)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8329
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8330
lemma SUM_DIFF: "ALL (x::'q_54364::type => hollight.real) (xa::'q_54364::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8331
   xb::'q_54364::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8332
   FINITE xa & SUBSET xb xa -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8333
   hollight.sum (DIFF xa xb) x =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8334
   real_sub (hollight.sum xa x) (hollight.sum xb x)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8335
  by (import hollight SUM_DIFF)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8336
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8337
lemma SUM_SUPPORT: "ALL (x::'q_54403::type => hollight.real) xa::'q_54403::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8338
   FINITE (support real_add x xa) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8339
   hollight.sum (support real_add x xa) x = hollight.sum xa x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8340
  by (import hollight SUM_SUPPORT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8341
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8342
lemma SUM_ADD: "ALL (f::'q_54449::type => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8343
   (g::'q_54449::type => hollight.real) s::'q_54449::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8344
   FINITE s -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8345
   hollight.sum s (%x::'q_54449::type. real_add (f x) (g x)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8346
   real_add (hollight.sum s f) (hollight.sum s g)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8347
  by (import hollight SUM_ADD)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8348
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8349
lemma SUM_CMUL: "ALL (f::'q_54487::type => hollight.real) (c::hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8350
   s::'q_54487::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8351
   FINITE s -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8352
   hollight.sum s (%x::'q_54487::type. real_mul c (f x)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8353
   real_mul c (hollight.sum s f)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8354
  by (import hollight SUM_CMUL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8355
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8356
lemma SUM_NEG: "ALL (x::'q_54530::type => hollight.real) xa::'q_54530::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8357
   FINITE xa -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8358
   hollight.sum xa (%xa::'q_54530::type. real_neg (x xa)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8359
   real_neg (hollight.sum xa x)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8360
  by (import hollight SUM_NEG)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8361
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8362
lemma SUM_SUB: "ALL (x::'q_54565::type => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8363
   (xa::'q_54565::type => hollight.real) xb::'q_54565::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8364
   FINITE xb -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8365
   hollight.sum xb (%xb::'q_54565::type. real_sub (x xb) (xa xb)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8366
   real_sub (hollight.sum xb x) (hollight.sum xb xa)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8367
  by (import hollight SUM_SUB)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8368
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8369
lemma SUM_LE: "ALL (x::'q_54607::type => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8370
   (xa::'q_54607::type => hollight.real) xb::'q_54607::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8371
   FINITE xb &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8372
   (ALL xc::'q_54607::type. IN xc xb --> real_le (x xc) (xa xc)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8373
   real_le (hollight.sum xb x) (hollight.sum xb xa)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8374
  by (import hollight SUM_LE)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8375
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8376
lemma SUM_LT: "ALL (f::'A::type => hollight.real) (g::'A::type => hollight.real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8377
   s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8378
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8379
   (ALL x::'A::type. IN x s --> real_le (f x) (g x)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8380
   (EX x::'A::type. IN x s & real_lt (f x) (g x)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8381
   real_lt (hollight.sum s f) (hollight.sum s g)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8382
  by (import hollight SUM_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8383
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8384
lemma SUM_LT_ALL: "ALL (f::'q_54729::type => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8385
   (g::'q_54729::type => hollight.real) s::'q_54729::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8386
   FINITE s &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8387
   s ~= EMPTY & (ALL x::'q_54729::type. IN x s --> real_lt (f x) (g x)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8388
   real_lt (hollight.sum s f) (hollight.sum s g)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8389
  by (import hollight SUM_LT_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8390
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8391
lemma SUM_EQ: "ALL (x::'q_54771::type => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8392
   (xa::'q_54771::type => hollight.real) xb::'q_54771::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8393
   FINITE xb & (ALL xc::'q_54771::type. IN xc xb --> x xc = xa xc) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8394
   hollight.sum xb x = hollight.sum xb xa"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8395
  by (import hollight SUM_EQ)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8396
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8397
lemma SUM_ABS: "ALL (f::'q_54830::type => hollight.real) s::'q_54830::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8398
   FINITE s -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8399
   real_le (real_abs (hollight.sum s f))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8400
    (hollight.sum s (%x::'q_54830::type. real_abs (f x)))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8401
  by (import hollight SUM_ABS)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8402
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8403
lemma SUM_CONST: "ALL (c::hollight.real) s::'q_54851::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8404
   FINITE s -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8405
   hollight.sum s (%n::'q_54851::type. c) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8406
   real_mul (real_of_num (CARD s)) c"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8407
  by (import hollight SUM_CONST)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8408
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8409
lemma SUM_EQ_0: "ALL (x::'A::type => hollight.real) xa::'A::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8410
   (ALL xb::'A::type. IN xb xa --> x xb = real_of_num 0) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8411
   hollight.sum xa x = real_of_num 0"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8412
  by (import hollight SUM_EQ_0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8413
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8414
lemma SUM_0: "ALL x::'A::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8415
   hollight.sum x (%n::'A::type. real_of_num 0) = real_of_num 0"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8416
  by (import hollight SUM_0)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8417
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8418
lemma SUM_POS_LE: "ALL (x::'q_54944::type => hollight.real) xa::'q_54944::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8419
   FINITE xa &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8420
   (ALL xb::'q_54944::type. IN xb xa --> real_le (real_of_num 0) (x xb)) -->
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8421
   real_le (real_of_num 0) (hollight.sum xa x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8422
  by (import hollight SUM_POS_LE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8423
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8424
lemma SUM_POS_BOUND: "ALL (f::'A::type => hollight.real) (b::hollight.real) x::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8425
   FINITE x &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8426
   (ALL xa::'A::type. IN xa x --> real_le (real_of_num 0) (f xa)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8427
   real_le (hollight.sum x f) b -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8428
   (ALL xa::'A::type. IN xa x --> real_le (f xa) b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8429
  by (import hollight SUM_POS_BOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8430
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8431
lemma SUM_POS_EQ_0: "ALL (x::'q_55091::type => hollight.real) xa::'q_55091::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8432
   FINITE xa &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8433
   (ALL xb::'q_55091::type. IN xb xa --> real_le (real_of_num 0) (x xb)) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8434
   hollight.sum xa x = real_of_num 0 -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8435
   (ALL xb::'q_55091::type. IN xb xa --> x xb = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8436
  by (import hollight SUM_POS_EQ_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8437
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8438
lemma SUM_SING: "ALL (x::'q_55113::type => hollight.real) xa::'q_55113::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8439
   hollight.sum (INSERT xa EMPTY) x = x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8440
  by (import hollight SUM_SING)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8441
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8442
lemma SUM_DELTA: "ALL (x::'A::type => bool) xa::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8443
   hollight.sum x
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8444
    (%x::'A::type. COND (x = xa) (b::hollight.real) (real_of_num 0)) =
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8445
   COND (IN xa x) b (real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8446
  by (import hollight SUM_DELTA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8447
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8448
lemma SUM_SWAP: "ALL (f::'A::type => 'B::type => hollight.real) (x::'A::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8449
   xa::'B::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8450
   FINITE x & FINITE xa -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8451
   hollight.sum x (%i::'A::type. hollight.sum xa (f i)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8452
   hollight.sum xa (%j::'B::type. hollight.sum x (%i::'A::type. f i j))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8453
  by (import hollight SUM_SWAP)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8454
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8455
lemma SUM_IMAGE: "ALL (x::'q_55257::type => 'q_55233::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8456
   (xa::'q_55233::type => hollight.real) xb::'q_55257::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8457
   FINITE xb &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8458
   (ALL (xa::'q_55257::type) y::'q_55257::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8459
       IN xa xb & IN y xb & x xa = x y --> xa = y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8460
   hollight.sum (IMAGE x xb) xa = hollight.sum xb (xa o x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8461
  by (import hollight SUM_IMAGE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8462
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8463
lemma SUM_SUPERSET: "ALL (f::'A::type => hollight.real) (u::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8464
   v::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8465
   FINITE u &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8466
   SUBSET u v &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8467
   (ALL x::'A::type. IN x v & ~ IN x u --> f x = real_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8468
   hollight.sum v f = hollight.sum u f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8469
  by (import hollight SUM_SUPERSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8470
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8471
lemma SUM_UNION_RZERO: "ALL (f::'A::type => hollight.real) (u::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8472
   v::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8473
   FINITE u &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8474
   (ALL x::'A::type. IN x v & ~ IN x u --> f x = real_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8475
   hollight.sum (hollight.UNION u v) f = hollight.sum u f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8476
  by (import hollight SUM_UNION_RZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8477
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8478
lemma SUM_UNION_LZERO: "ALL (f::'A::type => hollight.real) (u::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8479
   v::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8480
   FINITE v &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8481
   (ALL x::'A::type. IN x u & ~ IN x v --> f x = real_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8482
   hollight.sum (hollight.UNION u v) f = hollight.sum v f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8483
  by (import hollight SUM_UNION_LZERO)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8484
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8485
lemma SUM_RESTRICT: "ALL (f::'q_55484::type => hollight.real) s::'q_55484::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8486
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8487
   hollight.sum s
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8488
    (%x::'q_55484::type. COND (IN x s) (f x) (real_of_num 0)) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8489
   hollight.sum s f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8490
  by (import hollight SUM_RESTRICT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8491
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8492
lemma SUM_BOUND: "ALL (x::'A::type => bool) (xa::'A::type => hollight.real) xb::hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8493
   FINITE x & (ALL xc::'A::type. IN xc x --> real_le (xa xc) xb) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8494
   real_le (hollight.sum x xa) (real_mul (real_of_num (CARD x)) xb)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8495
  by (import hollight SUM_BOUND)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8496
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8497
lemma SUM_BOUND_GEN: "ALL (s::'A::type => bool) (t::'q_55543::type) b::hollight.real.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8498
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8499
   s ~= EMPTY &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8500
   (ALL x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8501
       IN x s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8502
       real_le ((f::'A::type => hollight.real) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8503
        (real_div b (real_of_num (CARD s)))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8504
   real_le (hollight.sum s f) b"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8505
  by (import hollight SUM_BOUND_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8506
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8507
lemma SUM_ABS_BOUND: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) b::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8508
   FINITE s & (ALL x::'A::type. IN x s --> real_le (real_abs (f x)) b) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8509
   real_le (real_abs (hollight.sum s f)) (real_mul (real_of_num (CARD s)) b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8510
  by (import hollight SUM_ABS_BOUND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8511
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8512
lemma SUM_BOUND_LT: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) b::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8513
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8514
   (ALL x::'A::type. IN x s --> real_le (f x) b) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8515
   (EX x::'A::type. IN x s & real_lt (f x) b) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8516
   real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8517
  by (import hollight SUM_BOUND_LT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8518
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8519
lemma SUM_BOUND_LT_ALL: "ALL (s::'q_55748::type => bool) (f::'q_55748::type => hollight.real)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8520
   b::hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8521
   FINITE s &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8522
   s ~= EMPTY & (ALL x::'q_55748::type. IN x s --> real_lt (f x) b) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8523
   real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8524
  by (import hollight SUM_BOUND_LT_ALL)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8525
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8526
lemma SUM_BOUND_LT_GEN: "ALL (s::'A::type => bool) (t::'q_55770::type) b::hollight.real.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8527
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8528
   s ~= EMPTY &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8529
   (ALL x::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8530
       IN x s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8531
       real_lt ((f::'A::type => hollight.real) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8532
        (real_div b (real_of_num (CARD s)))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8533
   real_lt (hollight.sum s f) b"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8534
  by (import hollight SUM_BOUND_LT_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8535
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8536
lemma SUM_UNION_EQ: "ALL (s::'q_55831::type => bool) (t::'q_55831::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8537
   u::'q_55831::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8538
   FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8539
   real_add (hollight.sum s (f::'q_55831::type => hollight.real))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8540
    (hollight.sum t f) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8541
   hollight.sum u f"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8542
  by (import hollight SUM_UNION_EQ)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8543
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8544
lemma SUM_EQ_SUPERSET: "ALL (f::'A::type => hollight.real) (s::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8545
   t::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8546
   FINITE t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8547
   SUBSET t s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8548
   (ALL x::'A::type. IN x t --> f x = (g::'A::type => hollight.real) x) &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8549
   (ALL x::'A::type. IN x s & ~ IN x t --> f x = real_of_num 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8550
   hollight.sum s f = hollight.sum t g"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8551
  by (import hollight SUM_EQ_SUPERSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8552
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8553
lemma SUM_RESTRICT_SET: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) r::'q_55944::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8554
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8555
   hollight.sum
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8556
    (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8557
      (%u::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8558
          EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8559
    f =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8560
   hollight.sum s (%x::'A::type. COND (P x) (f x) (real_of_num 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8561
  by (import hollight SUM_RESTRICT_SET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8562
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8563
lemma SUM_SUM_RESTRICT: "ALL (R::'q_56075::type => 'q_56074::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8564
   (f::'q_56075::type => 'q_56074::type => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8565
   (s::'q_56075::type => bool) t::'q_56074::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8566
   FINITE s & FINITE t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8567
   hollight.sum s
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8568
    (%x::'q_56075::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8569
        hollight.sum
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8570
         (GSPEC
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8571
           (%u::'q_56074::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8572
               EX y::'q_56074::type. SETSPEC u (IN y t & R x y) y))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8573
         (f x)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8574
   hollight.sum t
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8575
    (%y::'q_56074::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8576
        hollight.sum
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8577
         (GSPEC
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8578
           (%u::'q_56075::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8579
               EX x::'q_56075::type. SETSPEC u (IN x s & R x y) x))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8580
         (%x::'q_56075::type. f x y))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8581
  by (import hollight SUM_SUM_RESTRICT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8582
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8583
lemma CARD_EQ_SUM: "ALL x::'q_56096::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8584
   FINITE x -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8585
   real_of_num (CARD x) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8586
   hollight.sum x (%x::'q_56096::type. real_of_num (NUMERAL_BIT1 0))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8587
  by (import hollight CARD_EQ_SUM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8588
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8589
lemma SUM_MULTICOUNT_GEN: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8590
   (t::'B::type => bool) k::'B::type => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8591
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8592
   FINITE t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8593
   (ALL j::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8594
       IN j t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8595
       CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8596
        (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8597
          (%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8598
       k j) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8599
   hollight.sum s
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8600
    (%i::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8601
        real_of_num
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8602
         (CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8603
           (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8604
             (%u::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8605
                 EX j::'B::type. SETSPEC u (IN j t & R i j) j)))) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8606
   hollight.sum t (%i::'B::type. real_of_num (k i))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8607
  by (import hollight SUM_MULTICOUNT_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8608
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8609
lemma SUM_MULTICOUNT: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8610
   (t::'B::type => bool) k::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8611
   FINITE s &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8612
   FINITE t &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8613
   (ALL j::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8614
       IN j t -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8615
       CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8616
        (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8617
          (%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8618
       k) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8619
   hollight.sum s
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8620
    (%i::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8621
        real_of_num
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8622
         (CARD
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8623
           (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8624
             (%u::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8625
                 EX j::'B::type. SETSPEC u (IN j t & R i j) j)))) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8626
   real_of_num (k * CARD t)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8627
  by (import hollight SUM_MULTICOUNT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8628
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8629
lemma SUM_IMAGE_GEN: "ALL (f::'A::type => 'B::type) (g::'A::type => hollight.real)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8630
   s::'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8631
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8632
   hollight.sum s g =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8633
   hollight.sum (IMAGE f s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8634
    (%y::'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8635
        hollight.sum
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8636
         (GSPEC
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8637
           (%u::'A::type. EX x::'A::type. SETSPEC u (IN x s & f x = y) x))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8638
         g)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8639
  by (import hollight SUM_IMAGE_GEN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8640
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8641
lemma REAL_OF_NUM_SUM: "ALL (f::'q_56493::type => nat) s::'q_56493::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8642
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8643
   real_of_num (nsum s f) =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8644
   hollight.sum s (%x::'q_56493::type. real_of_num (f x))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8645
  by (import hollight REAL_OF_NUM_SUM)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8646
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8647
lemma SUM_SUBSET: "ALL (u::'A::type => bool) (v::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8648
   f::'A::type => hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8649
   FINITE u &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8650
   FINITE v &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8651
   (ALL x::'A::type. IN x (DIFF u v) --> real_le (f x) (real_of_num 0)) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8652
   (ALL x::'A::type. IN x (DIFF v u) --> real_le (real_of_num 0) (f x)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8653
   real_le (hollight.sum u f) (hollight.sum v f)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8654
  by (import hollight SUM_SUBSET)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8655
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8656
lemma SUM_SUBSET_SIMPLE: "ALL (u::'A::type => bool) (v::'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8657
   f::'A::type => hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8658
   FINITE v &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8659
   SUBSET u v &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8660
   (ALL x::'A::type. IN x (DIFF v u) --> real_le (real_of_num 0) (f x)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8661
   real_le (hollight.sum u f) (hollight.sum v f)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8662
  by (import hollight SUM_SUBSET_SIMPLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8663
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8664
lemma SUM_ADD_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8665
   hollight.sum (dotdot xb xc) (%i::nat. real_add (x i) (xa i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8666
   real_add (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8667
  by (import hollight SUM_ADD_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8668
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8669
lemma SUM_CMUL_NUMSEG: "ALL (x::nat => hollight.real) (xa::hollight.real) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8670
   hollight.sum (dotdot xb xc) (%i::nat. real_mul xa (x i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8671
   real_mul xa (hollight.sum (dotdot xb xc) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8672
  by (import hollight SUM_CMUL_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8673
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8674
lemma SUM_NEG_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8675
   hollight.sum (dotdot xa xb) (%i::nat. real_neg (x i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8676
   real_neg (hollight.sum (dotdot xa xb) x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8677
  by (import hollight SUM_NEG_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8678
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8679
lemma SUM_SUB_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8680
   hollight.sum (dotdot xb xc) (%i::nat. real_sub (x i) (xa i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8681
   real_sub (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8682
  by (import hollight SUM_SUB_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8683
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8684
lemma SUM_LE_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8685
   (ALL i::nat. <= xb i & <= i xc --> real_le (x i) (xa i)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8686
   real_le (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8687
  by (import hollight SUM_LE_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8688
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8689
lemma SUM_EQ_NUMSEG: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8690
   (ALL i::nat. <= m i & <= i n --> f i = g i) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8691
   hollight.sum (dotdot m n) f = hollight.sum (dotdot m n) g"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8692
  by (import hollight SUM_EQ_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8693
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8694
lemma SUM_ABS_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8695
   real_le (real_abs (hollight.sum (dotdot xa xb) x))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8696
    (hollight.sum (dotdot xa xb) (%i::nat. real_abs (x i)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8697
  by (import hollight SUM_ABS_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8698
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8699
lemma SUM_CONST_NUMSEG: "ALL (x::hollight.real) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8700
   hollight.sum (dotdot xa xb) (%n::nat. x) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8701
   real_mul (real_of_num (xb + NUMERAL_BIT1 0 - xa)) x"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8702
  by (import hollight SUM_CONST_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8703
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8704
lemma SUM_EQ_0_NUMSEG: "ALL (x::nat => hollight.real) xa::'q_57019::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8705
   (ALL i::nat. <= (m::nat) i & <= i (n::nat) --> x i = real_of_num 0) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8706
   hollight.sum (dotdot m n) x = real_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8707
  by (import hollight SUM_EQ_0_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8708
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8709
lemma SUM_TRIV_NUMSEG: "ALL (f::nat => hollight.real) (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8710
   < n m --> hollight.sum (dotdot m n) f = real_of_num 0"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8711
  by (import hollight SUM_TRIV_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8712
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8713
lemma SUM_POS_LE_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat => hollight.real.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8714
   (ALL p::nat. <= x p & <= p xa --> real_le (real_of_num 0) (xb p)) -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8715
   real_le (real_of_num 0) (hollight.sum (dotdot x xa) xb)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8716
  by (import hollight SUM_POS_LE_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8717
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8718
lemma SUM_POS_EQ_0_NUMSEG: "ALL (f::nat => hollight.real) (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8719
   (ALL p::nat. <= m p & <= p n --> real_le (real_of_num 0) (f p)) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8720
   hollight.sum (dotdot m n) f = real_of_num 0 -->
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8721
   (ALL p::nat. <= m p & <= p n --> f p = real_of_num 0)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8722
  by (import hollight SUM_POS_EQ_0_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8723
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8724
lemma SUM_SING_NUMSEG: "ALL (x::nat => hollight.real) xa::nat. hollight.sum (dotdot xa xa) x = x xa"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8725
  by (import hollight SUM_SING_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8726
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8727
lemma SUM_CLAUSES_NUMSEG: "(ALL x::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8728
    hollight.sum (dotdot x 0) (f::nat => hollight.real) =
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8729
    COND (x = 0) (f 0) (real_of_num 0)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8730
(ALL (x::nat) xa::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8731
    hollight.sum (dotdot x (Suc xa)) f =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8732
    COND (<= x (Suc xa))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8733
     (real_add (hollight.sum (dotdot x xa) f) (f (Suc xa)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8734
     (hollight.sum (dotdot x xa) f))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8735
  by (import hollight SUM_CLAUSES_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8736
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8737
lemma SUM_SWAP_NUMSEG: "ALL (a::nat) (b::nat) (c::nat) (d::nat) f::nat => nat => hollight.real.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8738
   hollight.sum (dotdot a b) (%i::nat. hollight.sum (dotdot c d) (f i)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8739
   hollight.sum (dotdot c d)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8740
    (%j::nat. hollight.sum (dotdot a b) (%i::nat. f i j))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8741
  by (import hollight SUM_SWAP_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8742
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8743
lemma SUM_ADD_SPLIT: "ALL (x::nat => hollight.real) (xa::nat) (xb::nat) xc::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8744
   <= xa (xb + NUMERAL_BIT1 0) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8745
   hollight.sum (dotdot xa (xb + xc)) x =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8746
   real_add (hollight.sum (dotdot xa xb) x)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8747
    (hollight.sum (dotdot (xb + NUMERAL_BIT1 0) (xb + xc)) x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8748
  by (import hollight SUM_ADD_SPLIT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8749
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8750
lemma SUM_OFFSET: "ALL (x::nat => hollight.real) (xa::nat) xb::nat.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8751
   hollight.sum (dotdot (xa + xb) ((n::nat) + xb)) x =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8752
   hollight.sum (dotdot xa n) (%i::nat. x (i + xb))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8753
  by (import hollight SUM_OFFSET)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8754
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8755
lemma SUM_OFFSET_0: "ALL (x::nat => hollight.real) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8756
   <= xa xb -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8757
   hollight.sum (dotdot xa xb) x =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8758
   hollight.sum (dotdot 0 (xb - xa)) (%i::nat. x (i + xa))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8759
  by (import hollight SUM_OFFSET_0)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8760
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8761
lemma SUM_CLAUSES_LEFT: "ALL (x::nat => hollight.real) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8762
   <= xa xb -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8763
   hollight.sum (dotdot xa xb) x =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8764
   real_add (x xa) (hollight.sum (dotdot (xa + NUMERAL_BIT1 0) xb) x)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8765
  by (import hollight SUM_CLAUSES_LEFT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8766
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8767
lemma SUM_CLAUSES_RIGHT: "ALL (f::nat => hollight.real) (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8768
   < 0 n & <= m n -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8769
   hollight.sum (dotdot m n) f =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8770
   real_add (hollight.sum (dotdot m (n - NUMERAL_BIT1 0)) f) (f n)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8771
  by (import hollight SUM_CLAUSES_RIGHT)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8772
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8773
lemma REAL_OF_NUM_SUM_NUMSEG: "ALL (x::nat => nat) (xa::nat) xb::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8774
   real_of_num (nsum (dotdot xa xb) x) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8775
   hollight.sum (dotdot xa xb) (%i::nat. real_of_num (x i))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8776
  by (import hollight REAL_OF_NUM_SUM_NUMSEG)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8777
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8778
lemma SUM_BIJECTION: "ALL (x::'A::type => hollight.real) (xa::'A::type => 'A::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8779
   xb::'A::type => bool.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8780
   FINITE xb &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8781
   (ALL x::'A::type. IN x xb --> IN (xa x) xb) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8782
   (ALL y::'A::type. IN y xb --> (EX! x::'A::type. IN x xb & xa x = y)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8783
   hollight.sum xb x = hollight.sum xb (x o xa)"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8784
  by (import hollight SUM_BIJECTION)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8785
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8786
lemma SUM_SUM_PRODUCT: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8787
   xb::'A::type => 'B::type => hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8788
   FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8789
   hollight.sum x (%x::'A::type. hollight.sum (xa x) (xb x)) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8790
   hollight.sum
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8791
    (GSPEC
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8792
      (%u::'A::type * 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8793
          EX (i::'A::type) j::'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8794
             SETSPEC u (IN i x & IN j (xa i)) (i, j)))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8795
    (GABS
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8796
      (%f::'A::type * 'B::type => hollight.real.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8797
          ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j)))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8798
  by (import hollight SUM_SUM_PRODUCT)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8799
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8800
lemma SUM_EQ_GENERAL: "ALL (x::'A::type => bool) (xa::'B::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8801
   (xb::'A::type => hollight.real) (xc::'B::type => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8802
   xd::'A::type => 'B::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8803
   FINITE x &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8804
   (ALL y::'B::type. IN y xa --> (EX! xa::'A::type. IN xa x & xd xa = y)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8805
   (ALL xe::'A::type. IN xe x --> IN (xd xe) xa & xc (xd xe) = xb xe) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8806
   hollight.sum x xb = hollight.sum xa xc"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8807
  by (import hollight SUM_EQ_GENERAL)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8808
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  8809
definition CASEWISE :: "(('q_57926 => 'q_57930) * ('q_57931 => 'q_57926 => 'q_57890)) hollight.list
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  8810
=> 'q_57931 => 'q_57930 => 'q_57890" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8811
  "CASEWISE ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8812
SOME CASEWISE::(('q_57926::type => 'q_57930::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8813
                ('q_57931::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8814
                 => 'q_57926::type => 'q_57890::type)) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8815
               => 'q_57931::type => 'q_57930::type => 'q_57890::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8816
   (ALL (f::'q_57931::type) x::'q_57930::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8817
       CASEWISE NIL f x = (SOME y::'q_57890::type. True)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8818
   (ALL (h::('q_57926::type => 'q_57930::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8819
            ('q_57931::type => 'q_57926::type => 'q_57890::type))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8820
       (t::(('q_57926::type => 'q_57930::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8821
            ('q_57931::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8822
             => 'q_57926::type => 'q_57890::type)) hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8823
       (f::'q_57931::type) x::'q_57930::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8824
       CASEWISE (CONS h t) f x =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8825
       COND (EX y::'q_57926::type. fst h y = x)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8826
        (snd h f (SOME y::'q_57926::type. fst h y = x)) (CASEWISE t f x))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8827
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8828
lemma DEF_CASEWISE: "CASEWISE =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8829
(SOME CASEWISE::(('q_57926::type => 'q_57930::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8830
                 ('q_57931::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8831
                  => 'q_57926::type => 'q_57890::type)) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8832
                => 'q_57931::type => 'q_57930::type => 'q_57890::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8833
    (ALL (f::'q_57931::type) x::'q_57930::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8834
        CASEWISE NIL f x = (SOME y::'q_57890::type. True)) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8835
    (ALL (h::('q_57926::type => 'q_57930::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8836
             ('q_57931::type => 'q_57926::type => 'q_57890::type))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8837
        (t::(('q_57926::type => 'q_57930::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8838
             ('q_57931::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8839
              => 'q_57926::type => 'q_57890::type)) hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8840
        (f::'q_57931::type) x::'q_57930::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8841
        CASEWISE (CONS h t) f x =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8842
        COND (EX y::'q_57926::type. fst h y = x)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8843
         (snd h f (SOME y::'q_57926::type. fst h y = x)) (CASEWISE t f x)))"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8844
  by (import hollight DEF_CASEWISE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8845
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8846
lemma CASEWISE: "(op &::bool => bool => bool)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8847
 ((op =::'q_57950::type => 'q_57950::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8848
   ((CASEWISE::(('q_57942::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8849
                ('q_57991::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8850
                 => 'q_57942::type => 'q_57950::type)) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8851
               => 'q_57991::type => 'q_57990::type => 'q_57950::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8852
     (NIL::(('q_57942::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8853
            ('q_57991::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8854
             => 'q_57942::type => 'q_57950::type)) hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8855
     (f::'q_57991::type) (x::'q_57990::type))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8856
   ((Eps::('q_57950::type => bool) => 'q_57950::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8857
     (%y::'q_57950::type. True::bool)))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8858
 ((op =::'q_57951::type => 'q_57951::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8859
   ((CASEWISE::(('q_57993::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8860
                ('q_57991::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8861
                 => 'q_57993::type => 'q_57951::type)) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8862
               => 'q_57991::type => 'q_57990::type => 'q_57951::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8863
     ((CONS::('q_57993::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8864
             ('q_57991::type => 'q_57993::type => 'q_57951::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8865
             => (('q_57993::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8866
                 ('q_57991::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8867
                  => 'q_57993::type => 'q_57951::type)) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8868
                => (('q_57993::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8869
                    ('q_57991::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8870
                     => 'q_57993::type => 'q_57951::type)) hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8871
       ((Pair::('q_57993::type => 'q_57990::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8872
               => ('q_57991::type => 'q_57993::type => 'q_57951::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8873
                  => ('q_57993::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8874
                     ('q_57991::type => 'q_57993::type => 'q_57951::type))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8875
         (s::'q_57993::type => 'q_57990::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8876
         (t::'q_57991::type => 'q_57993::type => 'q_57951::type))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8877
       (clauses::(('q_57993::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8878
                  ('q_57991::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8879
                   => 'q_57993::type => 'q_57951::type)) hollight.list))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8880
     f x)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8881
   ((COND::bool => 'q_57951::type => 'q_57951::type => 'q_57951::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8882
     ((Ex::('q_57993::type => bool) => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8883
       (%y::'q_57993::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8884
           (op =::'q_57990::type => 'q_57990::type => bool) (s y) x))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8885
     (t f ((Eps::('q_57993::type => bool) => 'q_57993::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8886
            (%y::'q_57993::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8887
                (op =::'q_57990::type => 'q_57990::type => bool) (s y) x)))
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8888
     ((CASEWISE::(('q_57993::type => 'q_57990::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8889
                  ('q_57991::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8890
                   => 'q_57993::type => 'q_57951::type)) hollight.list
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8891
                 => 'q_57991::type => 'q_57990::type => 'q_57951::type)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8892
       clauses f x)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8893
  by (import hollight CASEWISE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8894
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8895
lemma CASEWISE_CASES: "ALL (clauses::(('q_58085::type => 'q_58082::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8896
               ('q_58083::type
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8897
                => 'q_58085::type => 'q_58092::type)) hollight.list)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8898
   (c::'q_58083::type) x::'q_58082::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8899
   (EX (s::'q_58085::type => 'q_58082::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8900
       (t::'q_58083::type => 'q_58085::type => 'q_58092::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8901
       a::'q_58085::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8902
       MEM (s, t) clauses & s a = x & CASEWISE clauses c x = t c a) |
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8903
   ~ (EX (s::'q_58085::type => 'q_58082::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8904
         (t::'q_58083::type => 'q_58085::type => 'q_58092::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8905
         a::'q_58085::type. MEM (s, t) clauses & s a = x) &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8906
   CASEWISE clauses c x = (SOME y::'q_58092::type. True)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8907
  by (import hollight CASEWISE_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8908
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8909
lemma CASEWISE_WORKS: "ALL (x::(('P::type => 'A::type) *
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8910
         ('C::type => 'P::type => 'B::type)) hollight.list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8911
   xa::'C::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8912
   (ALL (s::'P::type => 'A::type) (t::'C::type => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8913
       (s'::'P::type => 'A::type) (t'::'C::type => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8914
       (xb::'P::type) y::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8915
       MEM (s, t) x & MEM (s', t') x & s xb = s' y -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8916
       t xa xb = t' xa y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8917
   ALL_list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8918
    (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8919
      (%f::('P::type => 'A::type) * ('C::type => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8920
           => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8921
          ALL (s::'P::type => 'A::type) t::'C::type => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8922
             GEQ (f (s, t))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8923
              (ALL xb::'P::type. CASEWISE x xa (s xb) = t xa xb)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8924
    x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8925
  by (import hollight CASEWISE_WORKS)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8926
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  8927
definition admissible :: "('q_58228 => 'q_58221 => bool)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8928
=> (('q_58228 => 'q_58224) => 'q_58234 => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8929
   => ('q_58234 => 'q_58221)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  8930
      => (('q_58228 => 'q_58224) => 'q_58234 => 'q_58229) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8931
  "admissible ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8932
%(u::'q_58228::type => 'q_58221::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8933
   (ua::('q_58228::type => 'q_58224::type) => 'q_58234::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8934
   (ub::'q_58234::type => 'q_58221::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8935
   uc::('q_58228::type => 'q_58224::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8936
       => 'q_58234::type => 'q_58229::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8937
   ALL (f::'q_58228::type => 'q_58224::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8938
      (g::'q_58228::type => 'q_58224::type) a::'q_58234::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8939
      ua f a &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8940
      ua g a & (ALL z::'q_58228::type. u z (ub a) --> f z = g z) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8941
      uc f a = uc g a"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8942
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8943
lemma DEF_admissible: "admissible =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8944
(%(u::'q_58228::type => 'q_58221::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8945
    (ua::('q_58228::type => 'q_58224::type) => 'q_58234::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8946
    (ub::'q_58234::type => 'q_58221::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8947
    uc::('q_58228::type => 'q_58224::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8948
        => 'q_58234::type => 'q_58229::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8949
    ALL (f::'q_58228::type => 'q_58224::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8950
       (g::'q_58228::type => 'q_58224::type) a::'q_58234::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8951
       ua f a &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8952
       ua g a & (ALL z::'q_58228::type. u z (ub a) --> f z = g z) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8953
       uc f a = uc g a)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8954
  by (import hollight DEF_admissible)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8955
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  8956
definition tailadmissible :: "('A => 'A => bool)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  8957
=> (('A => 'B) => 'P => bool)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  8958
   => ('P => 'A) => (('A => 'B) => 'P => 'B) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8959
  "tailadmissible ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8960
%(u::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8961
   (ua::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8962
   (ub::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8963
   uc::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8964
   EX (P::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8965
      (G::('A::type => 'B::type) => 'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8966
      H::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8967
      (ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8968
          P f a & u y (G f a) --> u y (ub a)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8969
      (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) a::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8970
          (ALL z::'A::type. u z (ub a) --> f z = g z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8971
          P f a = P g a & G f a = G g a & H f a = H g a) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8972
      (ALL (f::'A::type => 'B::type) a::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8973
          ua f a --> uc f a = COND (P f a) (f (G f a)) (H f a))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8974
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8975
lemma DEF_tailadmissible: "tailadmissible =
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8976
(%(u::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8977
    (ua::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8978
    (ub::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8979
    uc::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8980
    EX (P::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8981
       (G::('A::type => 'B::type) => 'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8982
       H::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8983
       (ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8984
           P f a & u y (G f a) --> u y (ub a)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8985
       (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) a::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8986
           (ALL z::'A::type. u z (ub a) --> f z = g z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8987
           P f a = P g a & G f a = G g a & H f a = H g a) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8988
       (ALL (f::'A::type => 'B::type) a::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8989
           ua f a --> uc f a = COND (P f a) (f (G f a)) (H f a)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8990
  by (import hollight DEF_tailadmissible)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8991
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  8992
definition superadmissible :: "('q_58378 => 'q_58378 => bool)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8993
=> (('q_58378 => 'q_58380) => 'q_58386 => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8994
   => ('q_58386 => 'q_58378)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34974
diff changeset
  8995
      => (('q_58378 => 'q_58380) => 'q_58386 => 'q_58380) => bool" where 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  8996
  "superadmissible ==
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8997
%(u::'q_58378::type => 'q_58378::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8998
   (ua::('q_58378::type => 'q_58380::type) => 'q_58386::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  8999
   (ub::'q_58386::type => 'q_58378::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9000
   uc::('q_58378::type => 'q_58380::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9001
       => 'q_58386::type => 'q_58380::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9002
   admissible u
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9003
    (%(f::'q_58378::type => 'q_58380::type) a::'q_58386::type. True) ub
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9004
    ua -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9005
   tailadmissible u ua ub uc"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9006
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9007
lemma DEF_superadmissible: "superadmissible =
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9008
(%(u::'q_58378::type => 'q_58378::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9009
    (ua::('q_58378::type => 'q_58380::type) => 'q_58386::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9010
    (ub::'q_58386::type => 'q_58378::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9011
    uc::('q_58378::type => 'q_58380::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9012
        => 'q_58386::type => 'q_58380::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9013
    admissible u
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9014
     (%(f::'q_58378::type => 'q_58380::type) a::'q_58386::type. True) ub
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9015
     ua -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9016
    tailadmissible u ua ub uc)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9017
  by (import hollight DEF_superadmissible)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9018
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9019
lemma SUPERADMISSIBLE_COND: "ALL (x::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9020
   (xa::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9021
   (xb::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9022
   (xc::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9023
   (xd::('A::type => 'B::type) => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9024
   xe::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9025
   admissible x xa xc xb &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9026
   superadmissible x
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9027
    (%(f::'A::type => 'B::type) x::'P::type. xa f x & xb f x) xc xd &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9028
   superadmissible x
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9029
    (%(f::'A::type => 'B::type) x::'P::type. xa f x & ~ xb f x) xc xe -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9030
   superadmissible x xa xc
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9031
    (%(f::'A::type => 'B::type) x::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9032
        COND (xb f x) (xd f x) (xe f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9033
  by (import hollight SUPERADMISSIBLE_COND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9034
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9035
lemma ADMISSIBLE_IMP_SUPERADMISSIBLE: "ALL (x::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9036
   (xa::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9037
   (xb::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9038
   xc::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9039
   admissible x xa xb xc --> superadmissible x xa xb xc"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9040
  by (import hollight ADMISSIBLE_IMP_SUPERADMISSIBLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9041
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9042
lemma TAIL_IMP_SUPERADMISSIBLE: "ALL (x::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9043
   (xa::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9044
   (xb::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9045
   xc::('A::type => 'B::type) => 'P::type => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9046
   admissible x xa xb xc &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9047
   (ALL (f::'A::type => 'B::type) a::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9048
       xa f a --> (ALL y::'A::type. x y (xc f a) --> x y (xb a))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9049
   superadmissible x xa xb
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9050
    (%(f::'A::type => 'B::type) x::'P::type. f (xc f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9051
  by (import hollight TAIL_IMP_SUPERADMISSIBLE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9052
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9053
lemma ADMISSIBLE_COND: "ALL (u_353::'A::type => 'q_58766::type => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9054
   (p::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9055
   (P::('A::type => 'B::type) => 'P::type => bool)
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9056
   (s::'P::type => 'q_58766::type)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9057
   (h::('A::type => 'B::type) => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9058
   k::('A::type => 'B::type) => 'P::type => 'B::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9059
   admissible u_353 p s P &
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9060
   admissible u_353 (%(f::'A::type => 'B::type) x::'P::type. p f x & P f x)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9061
    s h &
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9062
   admissible u_353
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9063
    (%(f::'A::type => 'B::type) x::'P::type. p f x & ~ P f x) s k -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9064
   admissible u_353 p s
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9065
    (%(f::'A::type => 'B::type) x::'P::type. COND (P f x) (h f x) (k f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9066
  by (import hollight ADMISSIBLE_COND)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9067
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9068
lemma ADMISSIBLE_CONST: "admissible (u_353::'q_58841::type => 'q_58840::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9069
 (p::('q_58841::type => 'q_58842::type) => 'q_58843::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9070
 (s::'q_58843::type => 'q_58840::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9071
 (%f::'q_58841::type => 'q_58842::type. c::'q_58843::type => 'q_58844::type)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9072
  by (import hollight ADMISSIBLE_CONST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9073
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9074
lemma ADMISSIBLE_COMB: "ALL (x::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9075
   (xa::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9076
   (xb::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9077
   (xc::('A::type => 'B::type) => 'P::type => 'C::type => 'D::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9078
   xd::('A::type => 'B::type) => 'P::type => 'C::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9079
   admissible x xa xb xc & admissible x xa xb xd -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9080
   admissible x xa xb
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9081
    (%(f::'A::type => 'B::type) x::'P::type. xc f x (xd f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9082
  by (import hollight ADMISSIBLE_COMB)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9083
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9084
lemma ADMISSIBLE_BASE: "ALL (x::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9085
   (xa::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9086
   (xb::'P::type => 'A::type) xc::'P::type => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9087
   (ALL (f::'A::type => 'B::type) a::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9088
       xa f a --> x (xc a) (xb a)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9089
   admissible x xa xb (%(f::'A::type => 'B::type) x::'P::type. f (xc x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9090
  by (import hollight ADMISSIBLE_BASE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9091
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9092
lemma ADMISSIBLE_NEST: "ALL (x::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9093
   (xa::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9094
   (xb::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9095
   xc::('A::type => 'B::type) => 'P::type => 'A::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9096
   admissible x xa xb xc &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9097
   (ALL (f::'A::type => 'B::type) a::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9098
       xa f a --> x (xc f a) (xb a)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9099
   admissible x xa xb (%(f::'A::type => 'B::type) x::'P::type. f (xc f x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9100
  by (import hollight ADMISSIBLE_NEST)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9101
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9102
lemma ADMISSIBLE_LAMBDA: "ALL (x::'A::type => 'A::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9103
   (xa::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9104
   (xb::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9105
   xc::'C::type => ('A::type => 'B::type) => 'P::type => 'D::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9106
   (ALL xd::'C::type. admissible x xa xb (xc xd)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9107
   admissible x xa xb
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9108
    (%(f::'A::type => 'B::type) (x::'P::type) u::'C::type. xc u f x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9109
  by (import hollight ADMISSIBLE_LAMBDA)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9110
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9111
lemma ADMISSIBLE_NSUM: "ALL (x::'B::type => 'A::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9112
   (xa::('B::type => 'C::type) => 'P::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9113
   (xb::'P::type => 'A::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9114
   (xc::('B::type => 'C::type) => 'P::type => nat => nat)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9115
   (xd::'P::type => nat) xe::'P::type => nat.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9116
   (ALL xf::nat.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9117
       admissible x
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9118
        (%(f::'B::type => 'C::type) x::'P::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9119
            <= (xd x) xf & <= xf (xe x) & xa f x)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9120
        xb (%(f::'B::type => 'C::type) x::'P::type. xc f x xf)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9121
   admissible x xa xb
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9122
    (%(f::'B::type => 'C::type) x::'P::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9123
        nsum (dotdot (xd x) (xe x)) (xc f x))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9124
  by (import hollight ADMISSIBLE_NSUM)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9125
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9126
lemma ADMISSIBLE_SUM: "ALL (x::'B::type => 'A::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9127
   (xa::('B::type => 'C::type) => 'P::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9128
   (xb::'P::type => 'A::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9129
   (xc::('B::type => 'C::type) => 'P::type => nat => hollight.real)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9130
   (xd::'P::type => nat) xe::'P::type => nat.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9131
   (ALL xf::nat.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9132
       admissible x
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9133
        (%(f::'B::type => 'C::type) x::'P::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9134
            <= (xd x) xf & <= xf (xe x) & xa f x)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9135
        xb (%(f::'B::type => 'C::type) x::'P::type. xc f x xf)) -->
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9136
   admissible x xa xb
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9137
    (%(f::'B::type => 'C::type) x::'P::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9138
        hollight.sum (dotdot (xd x) (xe x)) (xc f x))"
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9139
  by (import hollight ADMISSIBLE_SUM)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9140
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9141
lemma WF_REC_CASES: "ALL (u_353::'A::type => 'A::type => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9142
   clauses::(('P::type => 'A::type) *
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9143
             (('A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9144
              => 'P::type => 'B::type)) hollight.list.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9145
   WF u_353 &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9146
   ALL_list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9147
    (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9148
      (%f::('P::type => 'A::type) *
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9149
           (('A::type => 'B::type) => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9150
           => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9151
          ALL (s::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9152
             t::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9153
             GEQ (f (s, t))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9154
              (EX (P::('A::type => 'B::type) => 'P::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9155
                  (G::('A::type => 'B::type) => 'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9156
                  H::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9157
                  (ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9158
                      P f a & u_353 y (G f a) --> u_353 y (s a)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9159
                  (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9160
                      a::'P::type.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9161
                      (ALL z::'A::type. u_353 z (s a) --> f z = g z) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9162
                      P f a = P g a & G f a = G g a & H f a = H g a) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9163
                  (ALL (f::'A::type => 'B::type) a::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9164
                      t f a = COND (P f a) (f (G f a)) (H f a)))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9165
    clauses -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9166
   (EX f::'A::type => 'B::type. ALL x::'A::type. f x = CASEWISE clauses f x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9167
  by (import hollight WF_REC_CASES)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9168
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9169
lemma RECURSION_CASEWISE: "ALL clauses::(('P::type => 'A::type) *
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9170
              (('A::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9171
               => 'P::type => 'B::type)) hollight.list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9172
   (EX u::'A::type => 'A::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9173
       WF u &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9174
       ALL_list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9175
        (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9176
          (%f::('P::type => 'A::type) *
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9177
               (('A::type => 'B::type) => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9178
               => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9179
              ALL (s::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9180
                 t::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9181
                 GEQ (f (s, t))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9182
                  (tailadmissible u
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9183
                    (%(f::'A::type => 'B::type) a::'P::type. True) s t)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9184
        clauses) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9185
   (ALL (x::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9186
       (xa::('A::type => 'B::type) => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9187
       (xb::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9188
       (xc::('A::type => 'B::type) => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9189
       (xd::'A::type => 'B::type) (xe::'P::type) xf::'P::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9190
       MEM (x, xa) clauses & MEM (xb, xc) clauses -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9191
       x xe = xb xf --> xa xd xe = xc xd xf) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9192
   (EX f::'A::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9193
       ALL_list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9194
        (GABS
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9195
          (%fa::('P::type => 'A::type) *
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9196
                (('A::type => 'B::type) => 'P::type => 'B::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9197
                => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9198
              ALL (s::'P::type => 'A::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9199
                 t::('A::type => 'B::type) => 'P::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9200
                 GEQ (fa (s, t)) (ALL x::'P::type. f (s x) = t f x)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9201
        clauses)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9202
  by (import hollight RECURSION_CASEWISE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9203
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9204
lemma cth: "ALL (p1::'A::type => 'q_59947::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9205
   (p2::'q_59958::type => 'A::type => 'q_59952::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9206
   (p1'::'A::type => 'q_59947::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9207
   p2'::'q_59958::type => 'A::type => 'q_59952::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9208
   (ALL (c::'q_59958::type) (x::'A::type) y::'A::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9209
       p1 x = p1' y --> p2 c x = p2' c y) -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9210
   (ALL (c::'q_59958::type) (x::'A::type) y::'A::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9211
       p1' x = p1 y --> p2' c x = p2 c y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9212
  by (import hollight cth)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9213
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9214
lemma RECURSION_CASEWISE_PAIRWISE: "ALL x::(('q_59995::type => 'q_59975::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9215
        (('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9216
         => 'q_59995::type => 'q_59991::type)) hollight.list.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9217
   (EX u::'q_59975::type => 'q_59975::type => bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9218
       WF u &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9219
       ALL_list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9220
        (GABS
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9221
          (%f::('q_59995::type => 'q_59975::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9222
               (('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9223
                => 'q_59995::type => 'q_59991::type)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9224
               => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9225
              ALL (s::'q_59995::type => 'q_59975::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9226
                 t::('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9227
                    => 'q_59995::type => 'q_59991::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9228
                 GEQ (f (s, t))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9229
                  (tailadmissible u
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9230
                    (%(f::'q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9231
                        a::'q_59995::type. True)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9232
                    s t)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9233
        x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9234
   ALL_list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9235
    (GABS
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9236
      (%f::('q_59995::type => 'q_59975::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9237
           (('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9238
            => 'q_59995::type => 'q_59991::type)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9239
           => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9240
          ALL (a::'q_59995::type => 'q_59975::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9241
             b::('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9242
                => 'q_59995::type => 'q_59991::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9243
             GEQ (f (a, b))
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9244
              (ALL (c::'q_59975::type => 'q_59991::type) (x::'q_59995::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9245
                  y::'q_59995::type. a x = a y --> b c x = b c y)))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9246
    x &
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9247
   PAIRWISE
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9248
    (GABS
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9249
      (%f::('q_59995::type => 'q_59975::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9250
           (('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9251
            => 'q_59995::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9252
           => ('q_59995::type => 'q_59975::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9253
              (('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9254
               => 'q_59995::type => 'q_59991::type)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9255
              => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9256
          ALL (s::'q_59995::type => 'q_59975::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9257
             t::('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9258
                => 'q_59995::type => 'q_59991::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9259
             GEQ (f (s, t))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9260
              (GABS
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9261
                (%f::('q_59995::type => 'q_59975::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9262
                     (('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9263
                      => 'q_59995::type => 'q_59991::type)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9264
                     => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9265
                    ALL (s'::'q_59995::type => 'q_59975::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9266
                       t'::('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9267
                           => 'q_59995::type => 'q_59991::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9268
                       GEQ (f (s', t'))
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9269
                        (ALL (c::'q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9270
                            (x::'q_59995::type) y::'q_59995::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9271
                            s x = s' y --> t c x = t' c y)))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9272
    x -->
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9273
   (EX f::'q_59975::type => 'q_59991::type.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9274
       ALL_list
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9275
        (GABS
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9276
          (%fa::('q_59995::type => 'q_59975::type) *
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9277
                (('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9278
                 => 'q_59995::type => 'q_59991::type)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9279
                => bool.
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9280
              ALL (s::'q_59995::type => 'q_59975::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9281
                 t::('q_59975::type => 'q_59991::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9282
                    => 'q_59995::type => 'q_59991::type.
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9283
                 GEQ (fa (s, t)) (ALL x::'q_59995::type. f (s x) = t f x)))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9284
        x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9285
  by (import hollight RECURSION_CASEWISE_PAIRWISE)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9286
19093
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9287
lemma SUPERADMISSIBLE_T: "superadmissible (u_353::'q_60105::type => 'q_60105::type => bool)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9288
 (%(f::'q_60105::type => 'q_60107::type) x::'q_60111::type. True)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9289
 (s::'q_60111::type => 'q_60105::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9290
 (t::('q_60105::type => 'q_60107::type)
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9291
     => 'q_60111::type => 'q_60107::type) =
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9292
tailadmissible u_353
6d584f9d2021 use monomorphic sequences / scanners
obua
parents: 17915
diff changeset
  9293
 (%(f::'q_60105::type => 'q_60107::type) x::'q_60111::type. True) s t"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9294
  by (import hollight SUPERADMISSIBLE_T)
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9295
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9296
;end_setup
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9297
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9298
end
bd59bfd4bf37 fixed disambiguation problem
obua
parents:
diff changeset
  9299