Newer
Older
(** * Proofs for Section 5
everything here is based on the numbering in the extendend report
*)
Require Import GQM.Base.
Require Import GQM.GQM.
Require Import GQM.GQMDeduction.
Require Import GQM.GQMClassicalWithTactics.
Lemma lemma_A_1 : forall A B, |-- A ->> B -> |-- [A]A ->> [A]B.
move => A B HA.
have HB := (gqm_Anec _ HA).
eauto with GQMDB.
Qed.
(** *** Lemma A.2 *)
Lemma lemma_A_2_1 : forall A, |-- <E>A <<->> -![A]-!A.
Proof.
move => A.
unfold DiamondExists.
Qed.
Lemma lemma_A_2_2 : forall A, |-- <A>A <<->> -![E]-!A.
Proof.
move => A.
rewrite /DiamondAll /SquareExists /Exists /DiamondExists.
apply: gqm_dnegEquiv.
Qed.
Lemma lemma_A_2_3 : forall A, |-- [A]A <<->> -!<E>-!A.
Proof.
move => A.
hRewrite (gqm_dnegEquiv A) (fun X => [A]X <<->> -!<E>-!A).
hAndDestruct (lemma_A_2_1 (-! A)) => H1 H2.
- by apply: gqm_mp; first by apply: gqm_contrap3.
- by apply: gqm_mp; first by apply: gqm_contrap2.
Qed.
Lemma lemma_A_2_4 : forall A, |-- [E]A <<->> -!<A>-!A.
Proof.
move => A.
hAndDestruct (lemma_A_2_2 (-! A)) => ? ?.
hRewrite (gqm_dnegEquiv A) (fun X => [E]X <<->> -!(<A>-!A)).
- by apply: gqm_mp; first by apply: gqm_contrap3.
- by apply: gqm_mp; first by apply: gqm_contrap2.
Lemma lemma_A_2 : forall Q A, |-- (quant_to_neg_form Q)A <<->> -!((quant_to_form Q) (-! A)).
Proof.
case => /=.
- apply: lemma_A_2_1.
- apply: lemma_A_2_4.
- apply: lemma_A_2_2.
- apply: lemma_A_2_3.
Qed.
(** *** Proof of Lemma 5.2 *)
(*
Var 0 in A: p
Global Naming Context: 0 -> q
*)
Lemma lemma_A_3 : forall A, (|-- [A]A <<->> [A]A).
hSplit; first by apply: gqm_id.
(* (1) *)
have H1 := (gqm_Aglob (A._[ren_ (0 .: (+2))]) (Var 0)).
rewrite /subst_ /ren_ in H1.
asimpl in H1.
(* (2) *)
have Hr : (([A] A.[ren (0 .: (+2))]) = ([A]A).[ren (+1)]) by asimpl.
rewrite Hr in H1.
have H2 := (gqm_univgen _ _ H1).
asimpl in H2.
(* (3) *)
have H3 := (gqm_AE (A._[ren_ (+1)]) Bot).
rewrite /subst_ /ren_ in H3.
asimpl in H3.
(* (4) *)
have H4 := (lemma_A_1 _ _ H3).
(* (5) *)
apply: gqm_mp; last by apply: H2.
apply: gqm_mp; last by apply: H4.
apply: gqm_chain.
phi/A: 0->p, 1 -> r, 2 -> next element of global naming context after r
psi/B: 0->(local)r, 1 -> first element of global naming context after r
Global Naming Context: 0 -> r
.[ren (+1)] to express, that the global r cannot occur in the left term, because it is shadowed by the r in the binder
*)
Lemma lemma_A_4_1 : forall A B, |-- (([A]A.[B/]).[ren (+1)])->> [E]A.
(* have ? : (|-- (([A]((Var 3).[B/])).[ren (+1)])->> [E](Var 3)). { asimpl. } *)
rewrite /SquareExists /DiamondExists /All.
(* (1) is a noop here*)
(* (2) is a noop as well*)
(* (3) *)
have H3 := (gqm_AE (-])) B).
rewrite /subst_ /ren_ in H3.
(* (4) *)
asimpl in H3.
(* (5) *)
(* ._[ren (0 .: (+2))] to express the fact, that the global r does not occur in phi^p_psi, because it is shadowed by the r in the binder *)
have H5 := (gqm_Aglob (A._[B/]._[ren (0 .: (+2))]) (Var 0)).
rewrite /subst_ /ren_ /ids_ in H5.
(* 6 *)
apply: gqm_mp; last by apply: H5.
asimpl.
apply: gqm_mp; first by apply: gqm_chain.
(*
Global Naming Context antecedent: 0 -> q,...
Global Naming Context consequent: ...
*)
Lemma lemma_A_4_2 : forall X Phi A, |-- (X.[ren (+1)] & [A](Phi.[ren (+1)])) ->> A.[ren (+1)] -> |-- (X & [E]Phi) ->> A.
(* (1) *)
(* (2) *)
have H2 : (|-- (X.[ren (+1)] & -! A.[ren (+1)]) ->> -])) by apply: gqm_mp; first by apply: gqm_contrapAndR.
(* (3) *)
Fail have H3 : (|-- (X & -! A) ->> [A](-]))) by apply: gqm_univgen.
Lemma lemma_A_4_2_simple : forall Phi A, |-- [A](Phi.[ren (+1)]) ->> A.[ren (+1)] -> |-- [E]Phi ->> A.
Proof.
move => Phi A H.
apply: gqm_mp; last by apply: gqm_topI.
apply (tac_rewriteR (fun X => X) (gqm_andImpEquiv _ _ _)).
apply (tac_rewrite (fun X => X) (gqm_andImpEquiv _ _ _)).
Lemma lemma_A_5_1_1 : forall A B, |-- A ->> B -> |-- [A]A ->> [A]B.
Proof.
by apply: lemma_A_1.
Qed.
Lemma lemma_A_5_1_2 : forall A B, |-- A ->> B -> |-- [E]A ->> [E]B.
Proof.
move => A B Himp.
have H2 : (|-- [A](A.[ren (+1)]) ->> [A](B.[ren (+1)])). {
apply: lemma_A_5_1_1.
(* (3) *)
(* second ren (+1), because the second part does not refer to the global p*)
have H3 : (|-- [A](B.[ren (+1)]) ->> ([E]B).[ren (+1)]). {
(* Naming context for lemma_A_4_1:
* r -> 0, p -> 1, (Var 1) refers to p in global context
*)
have H' := (lemma_A_4_1 (B._[ren_ (0 .: (+3))]) (Var 1)).
rewrite /ren_ /subst_ /SquareExists /DiamondExists /All in H'.
asimpl in H'.
asimpl.
rewrite !ids_0_neg.
(* now we need to remove r->0 from the naming context *)
have H2' := (gqm_subst _ (ren_ (0 .: id)) H').
rewrite /ren_ in H2'.
by asimpl in H2'.
(* (4) *)
have H4 : (|-- [A](A.[ren (+1)]) ->> ([E]B).[ren (+1)]). {
apply: gqm_mp; last by apply: H2.
apply: gqm_mp; last by apply: H3.
by apply: gqm_chain.
}
(* (5) *)
by apply: lemma_A_4_2_simple.
Qed.
Lemma lemma_A_5_2_1 : forall A B, |-- A ->> B -> |-- <A>A ->> <A>B.
Proof.
move => A B Himp.
apply: (tac_rewrite (fun X => X ->> _) (lemma_A_2_2 _)).
apply: (tac_rewrite (fun X => _ ->> X) (lemma_A_2_2 _)).
apply: gqm_mp; first by apply: gqm_contrap.
apply: lemma_A_5_1_2.
by apply: gqm_mp; first by apply: gqm_contrap.
Qed.
Lemma lemma_A_5_2_2 : forall A B, |-- A ->> B -> |-- <E>A ->> <E>B.
Proof.
move => A B Himp.
apply: (tac_rewrite (fun X => X ->> _) (lemma_A_2_1 _)).
apply: (tac_rewrite (fun X => _ ->> X) (lemma_A_2_1 _)).
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
apply: gqm_mp; first by apply: gqm_contrap.
apply: lemma_A_5_1_1.
by apply: gqm_mp; first by apply: gqm_contrap.
Qed.
(** *** Lemma 5.3 *)
Lemma lemma_5_3 : forall Q A B, |-- A ->> B -> |-- (quant_to_form Q)A ->> (quant_to_form Q)B.
Proof.
case => /=.
- apply: lemma_A_5_1_1.
- apply: lemma_A_5_2_1.
- apply: lemma_A_5_1_2.
- apply: lemma_A_5_2_2.
Qed.
(** ** Lemma A.6 *)
Lemma lemma_A_6_1_1 : forall A, |-- [A]A ->> [E] A.
Proof.
move => A.
(* A with skipped global r *)
have H := (lemma_A_4_1 (A._[ren_ (0.: (+2)) ]) (Var 0)).
(* remove global r*)
have H2 := (gqm_subst _ (ren_ (0 .: id)) H).
rewrite /subst_ /ren_ in H2.
asimpl in H2.
by [].
Qed.
Lemma lemma_A_6_1_2 : forall A, |-- <A>A ->> <E> A.
Proof.
move => A.
apply: (tac_rewrite (fun X => X ->> _) (lemma_A_2_2 _)).
apply: (tac_rewrite (fun X => _ ->> X) (lemma_A_2_1 _)).
apply: gqm_mp; first by apply: gqm_contrap.
by apply: lemma_A_6_1_1.
Qed.
Lemma lemma_A_6_2_1 : forall A, |-- [E](A.[ren (+1)]) ->> [A](A.[ren (+1)]).
Proof.
move => A.
apply: lemma_A_4_2_simple.
asimpl.
by apply: gqm_id.
Qed.
Lemma lemma_A_6_2_2 : forall A, |-- <E>(A.[ren (+1)]) ->> <A>(A.[ren (+1)]).
Proof.
move => A.
apply: (tac_rewrite (fun X => X ->> _) (lemma_A_2_1 _)).
apply: (tac_rewrite (fun X => _ ->> X) (lemma_A_2_2 _)).
apply: gqm_mp; first by apply: gqm_contrap.
have H := (lemma_A_6_2_1 (-! A)).
by asimpl in H.
Lemma lemma_A_6_3_1 : forall A, |-- [A]A ->> <A>A.
Proof.
move => A.
rewrite /DiamondAll /Exists.
apply: lemma_A_1.
apply: gqm_mp; first by apply: gqm_contrap3.
have H := (gqm_AE (-! A._[ren_ (+1)]) (Var 0)).
rewrite /subst_ /ren_ in H.
by asimpl in H.
Qed.
Lemma lemma_A_6_3_2 : forall A, |-- [E]A ->> <E>A.
Proof.
move => A.
rewrite /SquareExists /DiamondExists.
(* other proof than in the paper *)
apply: gqm_mp; first by apply: gqm_contrap.
apply: (tac_rewrite (fun X => [A] -! A ->> [A] -! (|A| X)) (gqm_dnegEquiv _)).
Qed.
(** ** Lemma 5.4 / A.7 *)
Lemma lemma_A_7_1 : forall A B, |-- |A|(A ->> B) ->> (|A|A ->> |A|B).
Proof.
move => A B.
by apply: gqm_ADist.
Qed.
Lemma lemma_A_7_2_1 : forall A B, |-- |A|(A & B) <<->> (|A|A & |A| B).
Proof.
move => A B.
rewrite /And.
hSplit.
- apply: gqm_mp; last by apply: (lemma_A_7_1 (A ->> -! B) (Bot)).
apply: gqm_mp; first by apply: gqm_chain.
apply: (tac_rewrite (fun X => (|A| _ ->> X) ->> _) gqm_botAllEquiv).
apply: gqm_mp; first by apply: gqm_contrap.
(* TODO finish *)
Lemma lemma_A_7_2_2 : forall A B, |-- |E|(A ||| B) <<->> (|E|A ||| |E| B).
move => A B.
rewrite /Exists.
apply: (tac_rewriteR (fun X => X) (gqm_equivNegEquiv2 _ _)).
apply: (tac_rewrite (fun X => (|A| -! X) <<->> _) (gqm_demorganOr _ _)).
apply: (tac_rewriteR (fun X => (|A| X) <<->> _) (gqm_dnegEquiv _)).
apply: (tac_rewriteR (fun X => _ <<->> X) (gqm_demorganAnd _ _)).
apply: lemma_A_7_2_1.
Qed.
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
Lemma lemma_A_7_3_1 : forall A B, |-- A ->> B -> |-- |A|A ->> |A|B.
Proof.
move => A B H.
apply: lemma_A_5_1_1.
have -> :((A.[ren (+1)] ->> B.[ren (+1)]) = ((A ->> B).[ren (+1)])) by asimpl.
by apply: gqm_subst.
Qed.
Lemma lemma_A_7_3_2 : forall A B, |-- A ->> B -> |-- |E|A ->> |E|B.
Proof.
move => A B H.
rewrite /Exists.
apply: gqm_mp; first by apply: gqm_contrap.
apply: lemma_A_7_3_1.
by apply: gqm_mp; first by apply: gqm_contrap.
Qed.
Lemma lemma_A_7_4 : forall A, |-- |A|A ->> A.
Proof.
move => A.
have H := (gqm_AE (A._[ren_ (+1)]) (Var 0)).
rewrite /subst_ /ren_ in H.
by asimpl in H.
Qed.
Lemma lemma_A_7_5 : forall A, |-- A ->> |E|A.
Proof.
move => A.
rewrite /Exists.
apply: gqm_mp; first by apply: gqm_contrap3.
by apply: lemma_A_7_4.
Qed.
Lemma lemma_A_7_6 : forall A, |-- |E|A <<->> |A||E|A.
Proof.
hSplit; by [apply: gqm_Q5 | apply: lemma_A_7_4].
Qed.
Lemma lemma_A_7_7 : forall A, |-- |E||A|A <<->> |A|A.
Proof.
move => A.
rewrite /Exists.
(* TODO finish *)
Abort.
Lemma lemma_A_7_8_1 : forall A, |-- |A||A|A <<->> |A|A.
Proof.
move => A.
rewrite /Exists.
Abort.