Appendix 7
Coq Codes

A7.1. Some basic ontology and subtyping declarations

Definition CN:= Set.
Parameters Mary Helen: Woman.
Parameters John George: Man.
(** Common Nouns as Types *)
Parameters Phy Info Institution Woman Delegate : CN.
Parameters Man Object Surgeon Animal Human: CN.
(** Subtyping relations **)
Axiom sw: Surgeon -> Human. Coercion sw: Surgeon >-> Human. Axiom dh: Delegate -> Human. Coercion dh: Delegate >-> Human. Axiom mh: Man -> Human. Coercion mh: Man >-> Human.
Axiom wh: Woman -> Human. Coercion wh: Woman >-> Human.
Axiom ha: Human -> Animal. Coercion ha: Human>-> Animal.
Axiom ao: Animal -> Object. Coercion ao: Animal>-> Object.
Axiom bi: Bank -> Institution. Coercion bi: Bank >-> Institution.
(** Some quantifiers *)
Definition some:=
   fun A: CN=> fun P :A -> Prop=> exists x: A, P(x).
Definition all:=
   fun A: CN=> fun P: A -> Prop=> forall x: A, P(x).
Definition no:=
   fun A: CN=> fun P: A->Prop=> forall x: A, not(P(x)).

A7.2. Simple homonymy by overloading in coercive subtyping

(* Simple Homonymy by Overloading in Coercive Subtyping *) Set Implicit Arguments.
(* Unit type for "run" -- Onerun *)
Inductive Onerun: Set:= run.
Definition T1:= Human -> Prop.
Definition T2:= Human -> Institution -> Prop.
Parameter run1: T1. Parameter run2: T2.
Definition r1 (r: Onerun): T1:= run1. Coercion r1: Onerun >-> T1. Definition r2 (r: Onerun): T2:= run2. Coercion r2: Onerun >-> T2.
(*John runs quickly*)
Parameter quickly: forall (A: CN), (A -> Prop) -> (A -> Prop). Definition john_runs_quickly:= quickly (run:T1) John.
(*John runs a bank*)
Definition john_runs_a_bank:= exists b: Bank, (run: T2) John b.

A7.3. Intersective and subsective adjectives

(* Coq’s record types are Sigma-types *)
Parameter Irish: Human -> Prop.
Record Irishhuman: CN:= mkirishhuman
   {h:> Human; I: Irish h}.
Record Irishdelegate: CN:= mkirishdelegate
   {d:> Delegate; I2: Irish d}.
(*Skilful as a polymorphic type*)
Parameter skilful: forall A: CN, A -> Prop.
Record Skilfulman: CN:= mkskilfulman {m:> Man; S: skilful Man m}.
Record Skilfulhuman : CN:= mkskilfulhuman
   {h2:> Human; S2: skilful Human h2}.
Parameter walk: Human -> Prop.
Theorem delegate1:
   (some Irishdelegate) walk -> (some Delegate) walk. cbv. firstorder. Qed.
Theorem delegate2:
   (some Man) Irish -> (some Human) Irish.
cbv. firstorder. Qed.
Theorem skill1:
   (some Skilfulman) Irish -> (some Man) Irish.
cbv. firstorder. Qed.
Theorem skill2:
   (some Surgeon) (skilful Surgeon) -> (some Man)( skilful Man). cbv. firstorder. Abort all.

A7.4. Privative adjectives

(*Dealing with fake*)
Definition CN:= Set.
Parameter G_R: CN. Parameter G_F: CN.
Definition G:= (sum G_R G_F).
Definition fake_g (x: G):= match x with| inl _ => False | inr _ => True end.
Fixpoint real_g (x: G):= match x with| inl _ => True | inr _ => False end.
Parameter fake real: forall A: CN, A -> Prop.
Section FAKE.
Variable FAKE: fake G = fake_g. Variable REAL: real G = real_g.
Record fakegun: CN:= mkfakegun {g:> G; g1: fake G g}.
Record realgun: CN:= mkrealgun {r:> G; r1: real G r}.
(*That gun is either real or fake*)
Variable tg: G.
Theorem either:
   real G tg / fake G tg.
   replace (real(G)) with real_g. replace (fake G) with fake_g. cbv. destruct tg. left. trivial. right. trivial. Qed.

A7.5. Multidimensional adjectives

(*Dealing with multidimensional adjectives Health as an inductive type where the dimensions are enumerated. This is just an enumerated type*)
Definition Degree:= Set.
Inductive Health: Degree:= Heart|Blood|Cholesterol.
Parameter Healthy: Health -> Human -> Prop.
Definition sick:=fun y: Human => ~ (forall x : Health, Healthy x y).
Definition healthy:= fun y: Human => forall x: Health, Healthy x y.
Theorem HEALTHY:
   healthy John -> Healthy Heart John
   / Healthy Blood John
   / Healthy Cholesterol John.
   cbv. intros. split. apply H.
   split. apply H. apply H. Qed.
Theorem HEALTHY2:
   healthy John -> not (sick John).
   cbv. firstorder. Qed.
Theorem HEALTHY3:
   (exists x: Health, Healthy x John) -> healthy John. cbv. firstorder. Abort.
Theorem HEALTHY4:
   (exists x: Health, not (Healthy x John)) -> healthy John. cbv. firstorder. Abort.
Theorem HEALTHY5:
   (exists x: Health, not (Healthy x John)) -> sick John. cbv. firstorder. Qed.
(*Multidimensional noun Artist*)
Inductive Health: Degree:= a1|a2|a3.
Parameter DIM_CN : forall D: Degree, Human -> D -> Prop. Record Artist: Set:= mkartist
   {h:> Human; EI: forall a: art, (DIM_CN h a)}.

A7.6. Gradable adjectives

(*Degree is type of names of degrees
-- d: Degree corresponds to type D(d)*)
(*So, Degree is a Tarski universe!*)
(*Here is an example with three degrees*) Require Import Omega.
Inductive Degree: Set:= HEIGHT | AGE | IDIOCY.
Definition D (d: Degree):= nat.
Definition Height:= D(HEIGHT).
Definition Age:= D(AGE).
Definition Idiocy:= D(IDIOCY).
(*Universe CN_G of indexed CNs*)
Definition CN_G (_:Degree):= Set.
Parameter Human: CN_G(HEIGHT).
Parameter John Mary Kim: Human.
Parameter height: Human -> Height.
(*Type of physical objects indexed with a degree*) Parameter PHY : forall d: Degree, CN_G(d).
(*ADJ(D,A) of syntax of adjectives whose domain is A: CN_G(d)*) Parameter ADJ: forall d: Degree, CN_G(d)->Set.
Parameter TALL SHORT: ADJ HEIGHT Human.
Parameter IDIOTIC: ADJ IDIOCY Human.
Parameter ENORMOUS: forall d: Degree, ADJ d (PHY(d)).
(*Standard function*)
Parameter STND: forall d: Degree, forall A:CN_G(d), ADJ d A -> D(d).
(*semantics of tall, taller_than*)
Definition tall (h:Human):= ge (height h) (STND HEIGHT Human TALL).
Definition taller_than (h1:Human) (h2:Human) := gt (height h2) (height h1).
(*Some simple theorems*)
Theorem TALLER:
   taller_than Mary John / height Mary =
   170 -> gt (height John) 170.
   cbv. intro. omega. Qed.
Theorem trans:
   taller_than Mary John / taller_than Kim Mary ->
   taller_than Kim John.
   cbv. intro. omega. Qed.
(*Definition for idiot, enormous and enormous idiot*) Parameter IHuman: Idiocy -> CN_G(IDIOCY).
Definition Idiot:= sigT(fun x: Idiocy=> prod (IHuman x) (ge x (STND IDIOCY Human IDIOTIC))).
Definition enormous (d: Degree)(A: CN_G(d))(d1: D d):=
fun P: A => ge (d1) (STND d (PHY(d))(ENORMOUS d)).
Record enormousidiot: Set:= mkeidiot
   {h:> Idiot; EI: enormous IDIOCY
   (IHuman(projT1(h)))(projT1(h))(projT1(projT2(h)))
   / ge (STND IDIOCY (PHY(IDIOCY))(ENORMOUS IDIOCY)) (STND IDIOCY Human IDIOTIC)}.
(*From enormous idiot it follows that there exists an idiot such that their standard of idiocy is higher or equal than the standard for enormous idiots*)
Theorem ENORMOUS1:
   enormousidiot -> exists H: Idiot,
   projT1(H) >= STND IDIOCY (PHY IDIOCY) (ENORMOUS IDIOCY).
   cbv. firstorder. Qed.
(*From enormous idiot it follows that there exists an idiot such that their standard of idiocy is higher or equal to both the standard for enormous idiots and the standard for idiotic humans*)
Theorem ENORMOUS2:
   enormousidiot -> exists H: Idiot,
   projT1(H) >= STND IDIOCY (PHY IDIOCY) (ENORMOUS IDIOCY)
   / projT1(H) >= (STND IDIOCY Human IDIOTIC).
   cbv. firstorder. unfold Idiot in h0. exists h0. firstorder. unfold enormous in H. firstorder. elim h0. intros. destruct p. omega. Qed.
(*From enormous idiot it follows that there exists an idiot such that their standard of idiocy is higher or equal to the standard for enormous idiots and idiotic humans and also the standard for enormous idiots is higher than that for idiotic humans*)
Theorem ENORMOUS3:
   enormousidiot -> exists H: Idiot,
   projT1(H) >= STND IDIOCY (PHY IDIOCY) (ENORMOUS IDIOCY)
   / projT1(H) >= (STND IDIOCY Human IDIOTIC)
   / STND IDIOCY (PHY IDIOCY) (ENORMOUS IDIOCY)
   >= (STND IDIOCY Human IDIOTIC).
   cbv. firstorder. unfold Idiot in h0. exists h0. firstorder. unfold enormous in H. firstorder. elim h0. intros. destruct p. omega. Qed.

A7.7. Veridical adverbs

Definition VER_PROP (Q: Prop):= forall P: Prop, P -> Q. Parameter walk: Human -> Prop.
Definition fortunately:= VER_PROP.
Theorem VER1:
   fortunately (walk John) -> walk John.
   cbv. firstorder. apply H with (fortunately (walk John)).
   cbv. firstorder. Qed.
Set Implicit Arguments.
Definition VER_VP (A: CN) (Q: A -> Prop)(x: A):= forall P: A -> Prop, P x -> Q x.
Definition quickly:= VER_VP.
Theorem VER2:
   (quickly walk) John -> walk John.
   cbv. intro. apply H with (quickly walk).
   cbv. firstorder. Qed.
Theorem VER3:
   forall Q: Prop, VER_PROP Q -> Q.
   cbv. firstorder. apply H with (VER_PRO Q).
   cbv. assumption. Qed.
Theorem VER4 (P: Prop):
   forall Q: Prop, (P -> Q) -> (P -> VER_PROP Q).
   cbv. firstorder. Qed.

A7.8. Manner adverbs

Parameter Manner: Set.
Parameter Agent: CN.
Inductive Evt_m: Manner -> Type:= EVT_m : forall m: Manner, Evt_m m.
Inductive Evt_ma: Manner -> Agent -> Type:= EVT_ma : forall m: Manner, forall a: Agent, Evt_ma m a.
Parameter illegible_m:
   forall m: Manner, EVT_m m -> Prop. Parameter illegible_ma:
forall m: Manner, forall a: Agent, EVT_ma m a -> Prop.
Unset Implicit Arguments.
Parameter write: forall m: Manner, Human -> Evt_m m -> Prop. Set Implicit Arguments.
Definition illegibly_m:= fun (m: Manner)(A: CN) (P: A -> Evt_m m -> Prop) (x: A) (E: Evt_m m) =>
P x E / illegible_m m.
Parameter m: Manner.
Parameter u: Evt_m m.
Theorem MANNER:
   exists m: Manner, exists e: Evt_m m,
   illegibly_m (write m) John e -> illegible_m (m).
   exists m. exists u. cbv. firstorder. Qed.
Definition illegibly_a:= fun (m: Manner)(a: Agent)(A: CN) (P: A -> Evt_ma m a -> Prop) (x: A) (E: Evt_ma m a) => P x E / illegible_ma m a.

A7.9. Individuation

Require Import Setoid.
(*DomCN is the old CN, the universe of common nouns*) Definition DomCN:=Set.
(*Equiv is an equivalence relation on elements of DomCN*) Record Equiv (A:DomCN): Type:= mkEquiv
   {eq1:> A -> A -> Prop; _eq2: reflexive A(eq1) / symmetric A eq1 / transitive A eq1}.
Record Equiv_K (A:DomCN): Type:= mkEquiv_K
   {eq2:> A -> A -> Prop; _eq3: reflexive A(eq2) / symmetric A eq2}.
(*CN is a setoid, here expressed as a record, second projection is the equiv relation on the type*)
Record CN:= mkCN
   {D:> DomCN; E:Equiv D}.
Record CN_K:= mkCN_K
   {D2:> DomCN; E2:Equiv_K D2}.
(*Types in DomCN*)
Parameter Human Man: DomCN.
Parameter John: Human.
Axiom mh: Man->Human. Coercion mh: Man>->Human.
(*IC for Human*)
Parameter IC_Human: Equiv(Human).
(*This is the CN type in Capitals as a setoid*)
Definition HUMAN:= mkCN Human IC_Human.
(*Given the IC for man as being inherited from human*) Definition AIC_Man:= fun x: Man=>fun y: Man=> IC_Human(x)(y).
(*We prove the equivalence of AIC_Man*)
Theorem EQ:
   reflexive Man AIC_Man
   / symmetric Man AIC_Man
   / transitive Man AIC_Man.
   cbv. destruct IC_Human. destruct _eq4. destruct H0.
   split. intro. unfold reflexive in H. apply H. split.
   intro. unfold symmetric in H0. intuition. intro.
   unfold transitive in H1. intro. intro. intros.
   apply H1 with y. assumption. assumption. Qed.
(*A first definition of three, prety standard but takes into account the IC, B and B2 are the two respective projections*) Definition three_0:= fun A: CN => fun P: A.(D)->Prop=>
   exists x: A.(D),P(x) / exists y: A.(D), P(y)
   / exists z: A.(D), P(z) / not((A.(E))(x)(y))
   / not((A.(E))(y)(z)) / not((A.(E))(x)(z)).
(*Similarly for some*)
Definition some:=fun A: CN=>
   fun P: A.(D)->Prop=> exists x: A.(D), P(x).
(*Predicates are functions from the first CN projection to Prop*) Parameter walk: HUMAN.(D) -> Prop.
(*Defining the IC for man and then MAN*)
Definition IC_Man:= mkEquiv Man AIC_Man EQ.
Definition MAN:= mkCN Man IC_Man.
(*A proof that if three men walk, three humans walk*)
Theorem MANWALK:
   (three_0 MAN) walk-> (three_0 HUMAN) walk.
   cbv. intros. destruct H. destruct H. destruct H0.
   destruct H0. destruct H1. destruct H1. destruct H2.
   destruct H3. exists x. split. intuition. exists x0. split. intuition. exists x1. split. intuition. intuition. Qed.
(*Defining the dot type PhyInfo and declaring the coercions*) Parameters Phy: DomCN.
Parameter Info: DomCN.
Set Implicit Arguments.
Record PhyInfo: DomCN:= mkPhyInfo
   {p :> Phy; i :> Info}.
Parameter Book: DomCN.
Axiom bf: Book-> PhyInfo. Coercion bf: Book >-> PhyInfo.
(*IC for Phy, Info and PhyInfo*)
Parameter IC_Phy: Equiv(Phy).
Parameter IC_Info: Equiv(Info).
Definition PHY:= mkCN Phy IC_Phy.
Definition INFO:= mkCN Info IC_Info.
Definition AIC_PhyInfo := fun a1: (PhyInfo)=>fun b1: ( PhyInfo)=> IC_Phy(a1.(p))( b1.(p)) / IC_Info(a1.(i))( b1.(i)).
(*PhyInfo equiv is only reflexive and symmetric*)
Theorem EQ1:
   reflexive PhyInfo AIC_PhyInfo /
   symmetric PhyInfo AIC_PhyInfo.
   cbv. destruct IC_Phy. firstorder. destruct IC_Info.
   firstorder. Qed.
(*The _K partial equivalence is used for PhyInfo*) Definition IC_PhyInfo:= mkEquiv_K PhyInfo AIC_PhyInfo EQ1. Definition PHYINFO:= mkCN_K PhyInfo IC_PhyInfo.
(*Defining the IC criteria for book*)
Definition AIC_Book1:= fun x: Book=>fun y: Book=> IC_Phy(x)(y). Definition AIC_Book2:= fun x: Book=>fun y: Book=> IC_Info(x)(y).
Theorem EQ2:
   reflexive Book AIC_Book1
   / symmetric Book AIC_Book1
   / transitive Book AIC_Book1.
   cbv. destruct IC_Phy. destruct _eq4. destruct H0.
   split. intro. unfold reflexive in H. apply H. split. intro.
   unfold symmetric in H0. intuition. intro. unfold transitive in H1. intro. intro. intros. apply H1 with y. assumption.
   assumption. Qed.
Theorem EQ3:
   reflexive Book AIC_Book2
   / symmetric Book AIC_Book2
   / transitive Book AIC_Book2.
   cbv. destruct IC_Info. destruct _eq4. destruct H0. split. intro. unfold reflexive in H. apply H. split. intro. unfold symmetric in H0. intuition. intro. unfold transitive in H1. intro. intro. intros. apply H1 with y. assumption. assumption. Qed.
Definition IC_Book1:= mkEquiv Book AIC_Book1 EQ2.
Definition IC_Book2:= mkEquiv Book AIC_Book2 EQ3.
Definition BOOK1:= mkCN Book IC_Book1.
Definition BOOK2:= mkCN Book IC_Book2.
Parameter picked_up: Human -> Phy -> Prop. Parameter mastered: Human -> Info -> Prop.
Definition three:= fun A: DomCN => fun P: (PhyInfo) -> Prop=> exists x y z: Book, not(IC_PhyInfo((x))((y)))
   / not((IC_PhyInfo)((y))((z)))
   / not((IC_PhyInfo)((x))((z)))
   / P x / P y / P z.
Unset Implicit Arguments.
Definition and:= fun A: DomCN=> fun P: A->Prop=>fun Q: A->Prop=> fun x: A=> P(x) / Q(x).
(*Double distinctness case*)
Theorem copred_dd1:
   (three Book )(and PhyInfo(picked_up John)(mastered John)) -> (three_0 PHY) (picked_up John)
   / (three_0 INFO) (mastered John).
   cbv. intros. destruct IC_Phy. destruct IC_Info. firstorder. Qed.
(*Double distinctness with conclusion expanded from the start*) Theorem copred_dd2:
   (three Book )(and PhyInfo(picked_up John )(mastered John)) -> exists x y z: Book, not((PHY.(E))((x))((y)))
   / not((PHY.(E))((y))((z)))
   / not((PHY.(E))((x))((z)))
   / picked_up John (x)
   / picked_up John y
   / picked_up John z
   / not((INFO.(E))((x))((y)))
   / not((INFO.(E))((y))((z)))
   / ot((INFO.(E))((x))((z)))
   / mastered John (x)
   / mastered John y
   / mastered John z.
   cbv. firstorder. Qed.
Parameter Heavy: Phy -> Prop.
Record Heavybook: DomCN:= mkheavybook
   {b1:> Book; b2: Heavy b1}.
Definition AIC_HBook1 := fun x: Heavybook=>fun y: Heavybook=> IC_Book1(x)(y).
Theorem EQ4:
   reflexive Heavybook AIC_HBook1
   / symmetric Heavybook AIC_HBook1
   / transitive Heavybook AIC_HBook1.
   cbv. destruct IC_Phy. destruct _eq4. destruct H0. split. intro. unfold reflexive in H. apply H. split. intro. unfold symmetric in H0. intuition. intro. unfold transitive in H1. intro. intro. intros. apply H1 with y. assumption. assumption. Qed.
Definition IC_HBook1:= mkEquiv Heavybook AIC_HBook1 EQ4. Definition HBOOK1:= mkCN Heavybook IC_HBook1.
Definition AIC_HBook2:=fun x: Heavybook=>fun y:Heavybook=> IC_Book2(x)(y).
Theorem EQ5:
   reflexive Heavybook AIC_HBook2
   / symmetric Heavybook AIC_HBook2
   / transitive Heavybook AIC_HBook2.
   cbv. destruct IC_Info. destruct _eq4. destruct H0. split.
   intro. unfold reflexive in H. apply H. split. intro.
   unfold symmetric in H0. intuition. intro. unfold transitive in H1. intro. intro. intros. apply H1 with y. assumption. assumption. Qed.
Definition IC_HBook2:= mkEquiv Heavybook AIC_HBook2 EQ5. Definition HBOOK2:= mkCN Heavybook IC_HBook2.
Theorem HEAVYBOOK:
   (some HBOOK2) ( mastered John)-> (some INFO)( mastered John). firstorder. Qed.
Theorem HEAVYBOOK1:
   (some HBOOK1) (picked_up John)-> (some PHY)( picked_up John). firstorder. Qed.
Theorem HEAVYBOOK2:
   (three_0 HBOOK1) (picked_up John)->
   (three_0 PHY)( picked_up John).
   cbv. firstorder. Qed.
Theorem HEAVYBOOK3:
   (three_0 HBOOK2) ( mastered John)->
   (three_0 INFO)( mastered John).
   cbv. intros. firstorder. Qed.
(*More than one adjectives and individuation*)
Parameter Informative: Info->Prop.
Record Heavyinfobook: DomCN:= mkheavyinfobook
   {l4 :>Heavybook; _ : Informative l4}.
Definition AIC_HIBook1:= fun x:
   Heavyinfobook=> fun y: Heavyinfobook=> IC_HBook1(x)(y).
Theorem EQ6:
   reflexive Heavyinfobook AIC_HIBook1 /
   symmetric Heavyinfobook AIC_HIBook1 / transitive Heavyinfobook AIC_HIBook1.
   cbv. destruct IC_Phy. destruct _eq4. destruct H0. split.
   intro. unfold reflexive in H. apply H. split. intro. unfold symmetric
   in H0. intuition. intro. unfold transitive in H1. intro. intro. intros. apply H1 with y. assumption. assumption. Qed.
Definition IC_HIBook1:= mkEquiv Heavyinfobook AIC_HIBook1 EQ6. Definition HIBOOK1:= mkCN Heavyinfobook IC_HIBook1.
Definition AIC_HIBook2:= fun x: Heavyinfobook=
   >fun y: Heavyinfobook=> IC_HBook2(x)(y).
Theorem EQ7:
   reflexive Heavyinfobook AIC_HIBook2
   / symmetric Heavyinfobook AIC_HIBook2
   / transitive Heavyinfobook AIC_HIBook2.
   cbv. destruct IC_Info. destruct _eq4. destruct H0. split. intro. unfold reflexive in H. apply H. split. intro. unfold symmetric in H0. intuition. intro. unfold transitive in H1. intro. intro. intros. apply H1 with y. assumption. assumption. Qed.
Definition IC_HIBook2:= mkEquiv Heavyinfobook AIC_HIBook2 EQ7. Definition HIBOOK2:= mkCN Heavyinfobook IC_HIBook2.
Theorem HEAVYIBOOK2:
   (three_0 HBOOK1) (picked_up John)->
   (three_0 PHY)( picked_up John).
   cbv. firstorder. Qed.
Theorem HEAVYIBOOK3:
   (three_0 HIBOOK2) ( mastered John)->
   (three_0 INFO)( mastered John).
   cbv. intros. firstorder. Qed.
Theorem HEAVYIBOOK4:
   (three_0 HIBOOK1) (picked_up John)->
   (three_0 HBOOK1)( picked_up John).
   cbv. firstorder. Qed.
Theorem HEAVYIBOOK5:
   (three_0 HIBOOK2) ( mastered John) ->
   (three_0 HBOOK2)( mastered John).
   cbv. intros. firstorder. Qed.
..................Content has been hidden....................

You can't read the all page of ebook, please click here login for view all page.
Reset
18.117.216.36