Add noconf check

This commit is contained in:
Yiyun Liu 2025-02-28 16:39:10 -05:00
parent 7b5e9f2562
commit 3efca4160b

View file

@ -82,6 +82,26 @@ Definition isabs (a : PTm) :=
| _ => false | _ => false
end. end.
Definition tm_nonconf (a b : PTm) : bool :=
match a, b with
| PAbs _, _ => ishne b || isabs b
| _, PAbs _ => ishne a
| VarPTm _, VarPTm _ => true
| PPair _ _, _ => ishne b || ispair b
| _, PPair _ _ => ishne a
| PZero, PZero => true
| PSuc _, PSuc _ => true
| PApp _ _, PApp _ _ => ishne a && ishne b
| PProj _ _, PProj _ _ => ishne a && ishne b
| PInd _ _ _ _, PInd _ _ _ _ => ishne a && ishne b
| PNat, PNat => true
| PUniv _, PUniv _ => true
| PBind _ _ _, PBind _ _ _ => true
| _,_=> false
end.
Definition tm_conf (a b : PTm) := ~~ tm_nonconf a b.
Inductive eq_view : PTm -> PTm -> Type := Inductive eq_view : PTm -> PTm -> Type :=
| V_AbsAbs a b : | V_AbsAbs a b :
eq_view (PAbs a) (PAbs b) eq_view (PAbs a) (PAbs b)
@ -138,26 +158,6 @@ Equations tm_to_eq_view (a b : PTm) : eq_view a b :=
tm_to_eq_view (PBind p0 A0 B0) (PBind p1 A1 B1) := V_BindBind p0 A0 B0 p1 A1 B1; tm_to_eq_view (PBind p0 A0 B0) (PBind p1 A1 B1) := V_BindBind p0 A0 B0 p1 A1 B1;
tm_to_eq_view a b := V_Others a b. tm_to_eq_view a b := V_Others a b.
Definition tm_nonconf (a b : PTm) : bool :=
match a, b with
| PAbs _, _ => ishne b || isabs b
| _, PAbs _ => ishne a
| VarPTm _, VarPTm _ => true
| PPair _ _, _ => ishne b || ispair b
| _, PPair _ _ => ishne a
| PZero, PZero => true
| PSuc _, PSuc _ => true
| PApp _ _, PApp _ _ => ishne a && ishne b
| PProj _ _, PProj _ _ => ishne a && ishne b
| PInd _ _ _ _, PInd _ _ _ _ => ishne a && ishne b
| PNat, PNat => true
| PUniv _, PUniv _ => true
| PBind _ _ _, PBind _ _ _ => true
| _,_=> false
end.
Definition tm_conf (a b : PTm) := ~~ tm_nonconf a b.
Inductive algo_dom : PTm -> PTm -> Prop := Inductive algo_dom : PTm -> PTm -> Prop :=
| A_AbsAbs a b : | A_AbsAbs a b :
algo_dom_r a b -> algo_dom_r a b ->
@ -244,13 +244,6 @@ Inductive algo_dom : PTm -> PTm -> Prop :=
algo_dom_r c0 c1 -> algo_dom_r c0 c1 ->
algo_dom (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1) algo_dom (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
| A_Conflicting a b :
ishne a \/ ishf a ->
ishne b \/ ishf b ->
(* yes they are equivalent, but need both sides to make the rule reduce better *)
~ tm_nonconf a b ->
algo_dom a b
with algo_dom_r : PTm -> PTm -> Prop := with algo_dom_r : PTm -> PTm -> Prop :=
| A_NfNf a b : | A_NfNf a b :
algo_dom a b -> algo_dom a b ->
@ -345,7 +338,7 @@ Ltac check_equal_triv :=
intros;subst; intros;subst;
lazymatch goal with lazymatch goal with
(* | [h : algo_dom (VarPTm _) (PAbs _) |- _] => idtac *) (* | [h : algo_dom (VarPTm _) (PAbs _) |- _] => idtac *)
| [h : algo_dom _ _ |- _] => try (inversion h; subst;simpl in *; by first [done | exfalso; first [done|sfirstorder b:on]]) | [h : algo_dom _ _ |- _] => try (inversion h; by subst)
| _ => idtac | _ => idtac
end. end.
@ -356,20 +349,6 @@ Ltac solve_check_equal :=
| _ => idtac | _ => idtac
end]. end].
(* #[export,global] Obligation Tactic := idtac "fiewof". *)
Set Transparent Obligations.
(* Lemma algo_dom_abs_inv a b : *)
(* algo_dom (PAbs a) b -> ishne b \/ isabs b. *)
(* Proof. *)
(* inversion 1; subst=>//=. itauto. *)
(* itauto. *)
(* simpl in H2. *)
(* simpl in H2. move /negP in H2. move/norP in H2. *)
(* clear H0. left. *)
Equations check_equal (a b : PTm) (h : algo_dom a b) : Equations check_equal (a b : PTm) (h : algo_dom a b) :
bool by struct h := bool by struct h :=
check_equal a b h with tm_to_eq_view a b := check_equal a b h with tm_to_eq_view a b :=
@ -406,18 +385,9 @@ Equations check_equal (a b : PTm) (h : algo_dom a b) :
check_equal_r a b h with inspect (hred a) := check_equal_r a b h with inspect (hred a) :=
check_equal_r a b h (exist _ (Some a') k) := check_equal_r a' b _; check_equal_r a b h (exist _ (Some a') k) := check_equal_r a' b _;
check_equal_r a b h (exist _ None k) with inspect (hred b) := check_equal_r a b h (exist _ None k) with inspect (hred b) :=
| (exist _ None l) => check_equal a b _; | (exist _ None l) => tm_nonconf a b && check_equal a b _;
| (exist _ (Some b') l) => check_equal_r a b' _. | (exist _ (Some b') l) => check_equal_r a b' _.
Next Obligation.
inversion h; subst;simpl in *; try by first [done | exfalso; first [done|sfirstorder b:on]].
exfalso.
move /negP /norP in H1.
destruct H0 => //=. sfirstorder b:on.
Next Obligation. Next Obligation.
intros. intros.
destruct h. destruct h.