Newer
Older
require open personoj.encodings.lhol
require open personoj.encodings.pvs_cert
require open personoj.encodings.bool_hol
require open personoj.encodings.prenex
// definition neq {T: Set} (x y: El T) ≔ ¬ (x = y)
// symbol neq: χ (∀S (λt, scheme (t ~> t ~> bool)))
// rule neq _ $x $y ↪ ¬ ($x = $y)
// set infix left 2 "/=" ≔ neq
// set infix left 2 "≠" ≔ neq
/// Defined in builtins
//
// if_def
//
//
// boolean_props
// Slightly modified from the prelude
constant symbol bool_exclusive: Prf (neq {bool} false true)
Prf (∀ {bool} (λa, ((eq {bool} a false) ∨ (λ_, eq {bool} a true))))
theorem excluded_middle: Prf (∀ {bool} (λa, a ∨ (λ_, ¬ a)))
// PVS solves that kind of things thanks to the (bddsimp) tactic which uses an
// external C program
Prf (∀ {bool} (λa, ∀ {bool} (λb, eq {bool} (xor a b)
set prover "Alt-Ergo"
set prover_timeout 12
assume a b p
simpl
assume hxor
definition pred: ϕ (∀K (λ_, scheme_k {|set|})) ≔ λt, t ~> bool
definition PRED ≔ pred
definition predicate ≔ pred
definition PREDICATE ≔ pred
definition setof ≔ pred
definition SETOF ≔ pred
//
// exists1
//
//
// equality_props
//
: Prf (∀B (λt, ∀ (λx, ∀ {t} (λy, if true (λ_, x) (λ_, y) = x))))
: Prf (∀B (λt, ∀ (λx, ∀ {t} (λy, if false (λ_, x) (λ_, y) = y))))
: Prf (∀B (λt, ∀ {bool} (λb, ∀ (λx: El t, if b (λ_, x) (λ_, x) = x))))
constant symbol reflexivity_of_equals: Prf (∀B (λt, ∀ (λx: El t, x = x)))
: Prf (∀B (λt,
∀(λx: El t,
∀(λy: El t,
∀(λz: El t, (x = y) ∧ (λ_, y = z) ⊃ (λ_, x = z))))))
: Prf (∀B (λt, ∀ (λx: El t, (∀ (λy: El t, (x = y) ⊃ (λ_, y = x))))))
theorem eqind {t} x y: Prf (x = y) → Πp: El t → Bool, Prf (p y) → Prf (p x)
proof
assume t x y heq
apply symmetry_of_equals t x y heq
qed
set builtin "eqind" ≔ eqind
∀ (λx: El s,
∀ (λy: El s,
∀ (λf: El (s ~> t),
f (if a (λ_, x) (λ_, y))
= if a (λ_, f x) (λ_, f y))))))))
proof
admit
theorem lift_if2: