Add new definition of algo_dom
This commit is contained in:
parent
437c97455e
commit
849d19708e
4 changed files with 599 additions and 648 deletions
|
@ -11,26 +11,6 @@ Set Default Proof Mode "Classic".
|
|||
Require Import ssreflect ssrbool.
|
||||
From Hammer Require Import Tactics.
|
||||
|
||||
Definition tm_nonconf (a b : PTm) : bool :=
|
||||
match a, b with
|
||||
| PAbs _, _ => (~~ ishf b) || isabs b
|
||||
| _, PAbs _ => ~~ ishf a
|
||||
| VarPTm _, VarPTm _ => true
|
||||
| PPair _ _, _ => (~~ ishf b) || ispair b
|
||||
| _, PPair _ _ => ~~ ishf a
|
||||
| PZero, PZero => true
|
||||
| PSuc _, PSuc _ => true
|
||||
| PApp _ _, PApp _ _ => (~~ ishf a) && (~~ ishf b)
|
||||
| PProj _ _, PProj _ _ => (~~ ishf a) && (~~ ishf b)
|
||||
| PInd _ _ _ _, PInd _ _ _ _ => (~~ ishf a) && (~~ ishf 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 :=
|
||||
| V_AbsAbs a b :
|
||||
eq_view (PAbs a) (PAbs b)
|
||||
|
@ -102,121 +82,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 a b := V_Others a b _.
|
||||
|
||||
Inductive algo_dom : PTm -> PTm -> Prop :=
|
||||
| A_AbsAbs a b :
|
||||
algo_dom_r a b ->
|
||||
(* --------------------- *)
|
||||
algo_dom (PAbs a) (PAbs b)
|
||||
|
||||
| A_AbsNeu a u :
|
||||
ishne u ->
|
||||
algo_dom_r a (PApp (ren_PTm shift u) (VarPTm var_zero)) ->
|
||||
(* --------------------- *)
|
||||
algo_dom (PAbs a) u
|
||||
|
||||
| A_NeuAbs a u :
|
||||
ishne u ->
|
||||
algo_dom_r (PApp (ren_PTm shift u) (VarPTm var_zero)) a ->
|
||||
(* --------------------- *)
|
||||
algo_dom u (PAbs a)
|
||||
|
||||
| A_PairPair a0 a1 b0 b1 :
|
||||
algo_dom_r a0 a1 ->
|
||||
algo_dom_r b0 b1 ->
|
||||
(* ---------------------------- *)
|
||||
algo_dom (PPair a0 b0) (PPair a1 b1)
|
||||
|
||||
| A_PairNeu a0 a1 u :
|
||||
ishne u ->
|
||||
algo_dom_r a0 (PProj PL u) ->
|
||||
algo_dom_r a1 (PProj PR u) ->
|
||||
(* ----------------------- *)
|
||||
algo_dom (PPair a0 a1) u
|
||||
|
||||
| A_NeuPair a0 a1 u :
|
||||
ishne u ->
|
||||
algo_dom_r (PProj PL u) a0 ->
|
||||
algo_dom_r (PProj PR u) a1 ->
|
||||
(* ----------------------- *)
|
||||
algo_dom u (PPair a0 a1)
|
||||
|
||||
| A_ZeroZero :
|
||||
algo_dom PZero PZero
|
||||
|
||||
| A_SucSuc a0 a1 :
|
||||
algo_dom_r a0 a1 ->
|
||||
algo_dom (PSuc a0) (PSuc a1)
|
||||
|
||||
| A_UnivCong i j :
|
||||
(* -------------------------- *)
|
||||
algo_dom (PUniv i) (PUniv j)
|
||||
|
||||
| A_BindCong p0 p1 A0 A1 B0 B1 :
|
||||
algo_dom_r A0 A1 ->
|
||||
algo_dom_r B0 B1 ->
|
||||
(* ---------------------------- *)
|
||||
algo_dom (PBind p0 A0 B0) (PBind p1 A1 B1)
|
||||
|
||||
| A_NatCong :
|
||||
algo_dom PNat PNat
|
||||
|
||||
| A_VarCong i j :
|
||||
(* -------------------------- *)
|
||||
algo_dom (VarPTm i) (VarPTm j)
|
||||
|
||||
| A_ProjCong p0 p1 u0 u1 :
|
||||
ishne u0 ->
|
||||
ishne u1 ->
|
||||
algo_dom u0 u1 ->
|
||||
(* --------------------- *)
|
||||
algo_dom (PProj p0 u0) (PProj p1 u1)
|
||||
|
||||
| A_AppCong u0 u1 a0 a1 :
|
||||
ishne u0 ->
|
||||
ishne u1 ->
|
||||
algo_dom u0 u1 ->
|
||||
algo_dom_r a0 a1 ->
|
||||
(* ------------------------- *)
|
||||
algo_dom (PApp u0 a0) (PApp u1 a1)
|
||||
|
||||
| A_IndCong P0 P1 u0 u1 b0 b1 c0 c1 :
|
||||
ishne u0 ->
|
||||
ishne u1 ->
|
||||
algo_dom_r P0 P1 ->
|
||||
algo_dom u0 u1 ->
|
||||
algo_dom_r b0 b1 ->
|
||||
algo_dom_r c0 c1 ->
|
||||
algo_dom (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
|
||||
|
||||
| A_Conf a b :
|
||||
HRed.nf a ->
|
||||
HRed.nf b ->
|
||||
tm_conf a b ->
|
||||
algo_dom a b
|
||||
|
||||
with algo_dom_r : PTm -> PTm -> Prop :=
|
||||
| A_NfNf a b :
|
||||
algo_dom a b ->
|
||||
algo_dom_r a b
|
||||
|
||||
| A_HRedL a a' b :
|
||||
HRed.R a a' ->
|
||||
algo_dom_r a' b ->
|
||||
(* ----------------------- *)
|
||||
algo_dom_r a b
|
||||
|
||||
| A_HRedR a b b' :
|
||||
HRed.nf a ->
|
||||
HRed.R b b' ->
|
||||
algo_dom_r a b' ->
|
||||
(* ----------------------- *)
|
||||
algo_dom_r a b.
|
||||
|
||||
Scheme algo_ind := Induction for algo_dom Sort Prop
|
||||
with algor_ind := Induction for algo_dom_r Sort Prop.
|
||||
|
||||
Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
|
||||
|
||||
|
||||
(* Inductive salgo_dom : PTm -> PTm -> Prop := *)
|
||||
(* | S_UnivCong i j : *)
|
||||
|
@ -273,83 +138,6 @@ Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
|
|||
(* with algor_ind := Induction for salgo_dom_r Sort Prop. *)
|
||||
|
||||
|
||||
Lemma hf_no_hred (a b : PTm) :
|
||||
ishf a ->
|
||||
HRed.R a b ->
|
||||
False.
|
||||
Proof. hauto l:on inv:HRed.R. Qed.
|
||||
|
||||
Lemma hne_no_hred (a b : PTm) :
|
||||
ishne a ->
|
||||
HRed.R a b ->
|
||||
False.
|
||||
Proof. elim : a b => //=; hauto l:on inv:HRed.R. Qed.
|
||||
|
||||
Lemma algo_dom_no_hred (a b : PTm) :
|
||||
algo_dom a b ->
|
||||
HRed.nf a /\ HRed.nf b.
|
||||
Proof.
|
||||
induction 1 =>//=; try hauto inv:HRed.R use:hne_no_hred, hf_no_hred lq:on unfold:HRed.nf.
|
||||
Qed.
|
||||
|
||||
Derive Signature for algo_dom algo_dom_r.
|
||||
|
||||
Fixpoint hred (a : PTm) : option (PTm) :=
|
||||
match a with
|
||||
| VarPTm i => None
|
||||
| PAbs a => None
|
||||
| PApp (PAbs a) b => Some (subst_PTm (scons b VarPTm) a)
|
||||
| PApp a b =>
|
||||
match hred a with
|
||||
| Some a => Some (PApp a b)
|
||||
| None => None
|
||||
end
|
||||
| PPair a b => None
|
||||
| PProj p (PPair a b) => if p is PL then Some a else Some b
|
||||
| PProj p a =>
|
||||
match hred a with
|
||||
| Some a => Some (PProj p a)
|
||||
| None => None
|
||||
end
|
||||
| PUniv i => None
|
||||
| PBind p A B => None
|
||||
| PNat => None
|
||||
| PZero => None
|
||||
| PSuc a => None
|
||||
| PInd P PZero b c => Some b
|
||||
| PInd P (PSuc a) b c =>
|
||||
Some (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c)
|
||||
| PInd P a b c =>
|
||||
match hred a with
|
||||
| Some a => Some (PInd P a b c)
|
||||
| None => None
|
||||
end
|
||||
end.
|
||||
|
||||
Lemma hred_complete (a b : PTm) :
|
||||
HRed.R a b -> hred a = Some b.
|
||||
Proof.
|
||||
induction 1; hauto lq:on rew:off inv:HRed.R b:on.
|
||||
Qed.
|
||||
|
||||
Lemma hred_sound (a b : PTm):
|
||||
hred a = Some b -> HRed.R a b.
|
||||
Proof.
|
||||
elim : a b; hauto q:on dep:on ctrs:HRed.R.
|
||||
Qed.
|
||||
|
||||
Lemma hred_deter (a b0 b1 : PTm) :
|
||||
HRed.R a b0 -> HRed.R a b1 -> b0 = b1.
|
||||
Proof.
|
||||
move /hred_complete => + /hred_complete. congruence.
|
||||
Qed.
|
||||
|
||||
Definition fancy_hred (a : PTm) : HRed.nf a + {b | HRed.R a b}.
|
||||
destruct (hred a) eqn:eq.
|
||||
right. exists p. by apply hred_sound in eq.
|
||||
left. move => b /hred_complete. congruence.
|
||||
Defined.
|
||||
|
||||
Ltac2 destruct_algo () :=
|
||||
lazy_match! goal with
|
||||
| [h : algo_dom ?a ?b |- _ ] =>
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue