Skip to content

Commit 939e11d

Browse files
DmxLarcheymrhaandi
andauthored
Dealing with -notation-overridden compile option (#231)
* solved issues wrt overridden notations in the ILL/... sub-folders introducing modules for making notations more local. * revert the -notation-overridden compile option to active state * TM and SOL no warnings --------- Co-authored-by: Andrej Dudenhefner <[email protected]>
1 parent 86000d5 commit 939e11d

File tree

18 files changed

+127
-94
lines changed

18 files changed

+127
-94
lines changed

theories/ILL/CLL.v

Lines changed: 29 additions & 24 deletions
Original file line numberDiff line numberDiff line change
@@ -31,42 +31,47 @@ Inductive cll_form : Set :=
3131

3232
(* These notations replace the ILL notations *)
3333

34-
(* Variables *)
34+
Module CLL_notations.
3535

36-
Notation "£" := cll_var.
36+
(* Variables *)
3737

38-
(* Constants *)
38+
Notation "£" := cll_var.
3939

40-
Notation "⟙" := (cll_cst cll_top).
41-
Notation "⟘" := (cll_cst cll_bot).
42-
Notation "𝟙" := (cll_cst cll_1).
43-
Notation "𝟘" := (cll_cst cll_0).
40+
(* Constants *)
4441

45-
(* Unary connectives: linear negation and modalities *)
46-
(* ? cannot be used because it is reserved by Coq so we use ‽ instead *)
42+
Notation "⟙" := (cll_cst cll_top).
43+
Notation "⟘" := (cll_cst cll_bot).
44+
Notation "𝟙" := (cll_cst cll_1).
45+
Notation "𝟘" := (cll_cst cll_0).
4746

48-
Notation "'⊖' x" := (cll_una cll_neg x) (at level 50, format "⊖ x").
49-
Notation "'!' x" := (cll_una cll_bang x) (at level 52).
50-
Notation "'‽' x" := (cll_una cll_qmrk x) (at level 52).
47+
(* Unary connectives: linear negation and modalities *)
48+
(* ? cannot be used because it is reserved by Coq so we use ‽ instead *)
5149

52-
(* Binary connectives *)
50+
Notation "'⊖' x" := (cll_una cll_neg x) (at level 50, format "⊖ x").
51+
Notation "'!' x" := (cll_una cll_bang x) (at level 52).
52+
Notation "'‽' x" := (cll_una cll_qmrk x) (at level 52).
5353

54-
Infix "&" := (cll_bin cll_with) (at level 50).
55-
Infix "⅋" := (cll_bin cll_par) (at level 50).
56-
Infix "⊗" := (cll_bin cll_times) (at level 50).
57-
Infix "⊕" := (cll_bin cll_plus) (at level 50).
58-
Infix "⊸" := (cll_bin cll_limp) (at level 51, right associativity).
54+
(* Binary connectives *)
5955

60-
(* Modalities iterated over lists *)
56+
Infix "&" := (cll_bin cll_with) (at level 50).
57+
Infix "⅋" := (cll_bin cll_par) (at level 50).
58+
Infix "⊗" := (cll_bin cll_times) (at level 50).
59+
Infix "⊕" := (cll_bin cll_plus) (at level 50).
60+
Infix "⊸" := (cll_bin cll_limp) (at level 51, right associativity).
6161

62-
Notation "‼ x" := (map (cll_una cll_bang) x) (at level 60).
63-
Notation "⁇ x" := (map (cll_una cll_qmrk) x) (at level 60).
62+
(* Modalities iterated over lists *)
6463

65-
(* The empty list *)
64+
Notation "‼ x" := (map (cll_una cll_bang) x) (at level 60).
65+
Notation "⁇ x" := (map (cll_una cll_qmrk) x) (at level 60).
66+
67+
End CLL_notations.
6668

67-
Notation "∅" := nil.
69+
Import CLL_notations.
70+
71+
(* The empty list *)
6872

69-
Local Reserved Notation "Γ ⊢ Δ" (at level 70, no associativity).
73+
#[local] Notation "∅" := nil.
74+
#[local] Reserved Notation "Γ ⊢ Δ" (at level 70, no associativity).
7075

7176
Section S_cll_restr_without_cut.
7277

theories/ILL/EILL.v

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -13,6 +13,8 @@ From Stdlib Require Import List Permutation.
1313

1414
From Undecidability.ILL Require Import ILL.
1515

16+
Import ILL_notations.
17+
1618
Set Implicit Arguments.
1719

1820
Local Infix "~p" := (@Permutation _) (at level 70).

theories/ILL/ILL.v

Lines changed: 18 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -35,24 +35,30 @@ Inductive ill_form : Set :=
3535

3636
(* Symbols for cut&paste ⟙ ⟘ 𝝐 ﹠ ⊗ ⊕ ⊸ ! ‼ ∅ ⊢ *)
3737

38-
Notation "⟙" := (ill_cst ill_top).
39-
Notation "⟘" := (ill_cst ill_bot).
40-
Notation "𝟙" := (ill_cst ill_1).
38+
Module ILL_notations.
4139

42-
Infix "&" := (ill_bin ill_with) (at level 50).
43-
Infix "⊗" := (ill_bin ill_times) (at level 50).
44-
Infix "⊕" := (ill_bin ill_plus) (at level 50).
45-
Infix "⊸" := (ill_bin ill_limp) (at level 51, right associativity).
40+
Notation "⟙" := (ill_cst ill_top).
41+
Notation "⟘" := (ill_cst ill_bot).
42+
Notation "𝟙" := (ill_cst ill_1).
4643

47-
Notation "'!' x" := (ill_ban x) (at level 52).
44+
Infix "&" := (ill_bin ill_with) (at level 50).
45+
Infix "⊗" := (ill_bin ill_times) (at level 50).
46+
Infix "⊕" := (ill_bin ill_plus) (at level 50).
47+
Infix "⊸" := (ill_bin ill_limp) (at level 51, right associativity).
4848

49-
Notation "£" := ill_var.
49+
Notation "'!' x" := (ill_ban x) (at level 52).
5050

51-
Notation "‼ x" := (map ill_ban x) (at level 60).
51+
Notation "£" := ill_var.
5252

53-
Notation "" := nil (only parsing).
53+
Notation "‼ x" := (map ill_ban x) (at level 60).
5454

55-
Reserved Notation "l '⊢' x" (at level 70, no associativity).
55+
Notation "∅" := nil (only parsing).
56+
57+
End ILL_notations.
58+
59+
Import ILL_notations.
60+
61+
#[local] Reserved Notation "l '⊢' x" (at level 70, no associativity).
5662

5763
Section S_ill_restr_without_cut.
5864

theories/ILL/IMSELL.v

Lines changed: 18 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -17,9 +17,9 @@ Set Implicit Arguments.
1717

1818
Local Infix "~p" := (@Permutation _) (at level 70).
1919

20-
Reserved Notation "A ⊸ B" (at level 51, right associativity, format "A ⊸ B").
21-
Reserved Notation "'![' m ']' x" (at level 52, format "![ m ] x").
22-
Reserved Notation "‼ x" (at level 60, format "‼ x").
20+
#[local] Reserved Notation "A ⊸ B" (at level 51, right associativity, format "A ⊸ B").
21+
#[local] Reserved Notation "'![' m ']' x" (at level 52, format "![ m ] x").
22+
#[local] Reserved Notation "‼ x" (at level 60, format "‼ x").
2323

2424
Section IMSELL.
2525

@@ -85,10 +85,21 @@ Section IMSELL.
8585

8686
End IMSELL.
8787

88-
Infix "⊸" := imsell_imp.
89-
Notation "![ m ] x" := (imsell_ban m x).
90-
Notation "£" := imsell_var.
91-
Notation "‼ Γ" := (imsell_lban Γ).
88+
Arguments imsell_var {_ _}.
89+
Arguments imsell_imp {_ _}.
90+
Arguments imsell_ban {_ _}.
91+
Arguments imsell_lban {_ _}.
92+
93+
Module IMSELL_notations.
94+
95+
Infix "⊸" := imsell_imp.
96+
Notation "![ m ] x" := (imsell_ban m x).
97+
Notation "£" := imsell_var.
98+
Notation "‼ Γ" := (imsell_lban Γ).
99+
100+
End IMSELL_notations.
101+
102+
Import IMSELL_notations.
92103

93104
(* An IMSELL signature is a type of modalities pre-ordered
94105
and an upper-closed subset of exponentials *)

theories/ILL/Ll/eill.v

Lines changed: 5 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -14,15 +14,14 @@ From Undecidability.Shared.Libs.DLW
1414

1515
From Undecidability.ILL Require Import ILL EILL ill.
1616

17-
Set Implicit Arguments.
17+
Import ILL_notations.
1818

19-
Local Infix "~p" := (@Permutation _) (at level 70).
19+
Set Implicit Arguments.
2020

2121
(* Symbols for cut&paste ⟙ ⟘ 𝝐 ﹠ ⊗ ⊕ ⊸ ❗ ‼ ∅ ⊢ ⟦ ⟧ Γ Δ Σ *)
2222

23-
Notation "⦑ c ⦒" := (eill_cmd_map c) (at level 0).
24-
25-
Notation "Σ ; Γ ⊦ u" := (G_eill Σ Γ u) (at level 70, no associativity).
23+
#[local] Notation "⦑ c ⦒" := (eill_cmd_map c) (at level 0).
24+
#[local] Notation "Σ ; Γ ⊦ u" := (G_eill Σ Γ u) (at level 70, no associativity).
2625

2726
Theorem g_eill_mono_Si Σ Σ' Γ u : incl Σ Σ' -> Σ; Γ ⊦ u -> Σ'; Γ ⊦ u.
2827
Proof.
@@ -47,7 +46,7 @@ Qed.
4746
the cut-free (!,&,-o) fragment
4847
*)
4948

50-
Notation "Γ ⊢ A" := (S_ill_restr Γ A) (at level 70, no associativity).
49+
#[local] Notation "Γ ⊢ A" := (S_ill_restr Γ A) (at level 70, no associativity).
5150

5251
Theorem G_eill_sound Σ Γ p : Σ; Γ ⊦ p -> map (fun c => !⦑c⦒) Σ ++ map £ Γ ⊢ £ p.
5352
Proof.

theories/ILL/Ll/eill_mm.v

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -18,9 +18,12 @@ From Undecidability.MinskyMachines.MM
1818
From Undecidability.ILL
1919
Require Import ILL EILL ill eill.
2020

21+
Import ILL_notations.
22+
2123
Set Implicit Arguments.
2224

23-
Local Infix "~p" := (@Permutation _) (at level 70).
25+
#[local] Notation "⦑ c ⦒" := (eill_cmd_map c) (at level 0).
26+
#[local] Notation "Σ ; Γ ⊦ u" := (G_eill Σ Γ u) (at level 70, no associativity).
2427

2528
(* ** MM reduces to eILL *)
2629

theories/ILL/Ll/ill.v

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -15,6 +15,8 @@ From Undecidability.Shared.Libs.DLW
1515
From Undecidability.ILL
1616
Require Import ILL.
1717

18+
Import ILL_notations.
19+
1820
Set Implicit Arguments.
1921

2022
Section obvious_links_between_fragments.

theories/ILL/Ll/ill_cll.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -127,8 +127,8 @@ Local Hint Resolve ill_cll_ill : core.
127127
Fact ill_cll_ill_list Γ : ⟪⟦Γ⟧⟫ = Γ.
128128
Proof. induction Γ; simpl; f_equal; auto. Qed.
129129

130-
Fact ill_cll_lbang Γ : ⟦map ill_ban Γ⟧ = ⟦Γ⟧.
130+
Fact ill_cll_lbang Γ : ⟦map ill_ban Γ⟧ = map (cll_una cll_bang) ⟦Γ⟧.
131131
Proof. induction Γ; simpl; f_equal; auto. Qed.
132132

133-
Fact cll_ill_lbang Γ : ⟪Γ⟫ = map ill_ban ⟪Γ⟫.
133+
Fact cll_ill_lbang Γ : ⟪map (cll_una cll_bang) Γ⟫ = map ill_ban ⟪Γ⟫.
134134
Proof. induction Γ; simpl; f_equal; auto. Qed.

theories/ILL/Ll/ill_cll_restr.v

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -37,6 +37,7 @@ Local Infix "~p" := (@Permutation _) (at level 70).
3737
3838
derived from the work of H. Schellinx JLC 91 *)
3939

40+
#[local] Notation "∅" := nil.
4041

4142
Section S_ill_cll_restr.
4243

theories/ILL/Ll/imsell.v

Lines changed: 6 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -12,25 +12,24 @@ From Stdlib Require Import List Permutation.
1212
From Undecidability.ILL
1313
Require Import IMSELL.
1414

15+
Import IMSELL_notations.
16+
1517
From Undecidability.Shared.Libs.DLW
1618
Require Import utils pos vec.
1719

1820
Set Implicit Arguments.
1921

2022
(* * Intuionistic Multiplicative Linear Logic with several exponentials and modabilities *)
2123

22-
Local Infix "~p" := (@Permutation _) (at level 70).
24+
(* Local Infix "~p" := (@Permutation _) (at level 70). *)
2325
Local Notation "X ⊆ Y" := (forall a, X a -> Y a : Prop) (at level 70).
2426
Local Infix "∊" := In (at level 70).
2527

26-
Local Reserved Notation "'⟦' A '⟧'" (at level 0, format "⟦ A ⟧").
27-
2828
Section IMSELL.
2929

3030
Variable bang : Type.
3131

32-
Notation "£" := (@imsell_var _ _).
33-
Notation "‼ l" := (@imsell_lban nat bang l).
32+
Implicit Type (Σ : list (bang * imsell_form nat bang)).
3433

3534
Fact imsell_lban_perm Σ Γ : Σ ~p Γ -> ‼Σ ~p ‼Γ.
3635
Proof. apply Permutation_map. Qed.
@@ -102,6 +101,8 @@ Section IMSELL.
102101
103102
*)
104103

104+
Reserved Notation "'⟦' A '⟧'" (at level 0, format "⟦ A ⟧").
105+
105106
Section Trivial_Phase_semantics.
106107

107108
(* We only consider the monoids nat^n *)

0 commit comments

Comments
 (0)